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(-) 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(-) 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(-) 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 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 --- driver/VericertDriver.ml | 11 ++++--- src/Compiler.v | 1 + src/VericertClflags.ml | 1 + src/extraction/Extraction.v | 1 + src/hls/printRTLPar.ml | 74 +++++++++++++++++++++++++++++++++++++++++++++ 5 files changed, 84 insertions(+), 4 deletions(-) create mode 100644 src/hls/printRTLPar.ml diff --git a/driver/VericertDriver.ml b/driver/VericertDriver.ml index d2c301f..4c1e2f9 100644 --- a/driver/VericertDriver.ml +++ b/driver/VericertDriver.ml @@ -66,6 +66,7 @@ let compile_c_file sourcename ifile ofile = set_dest Vericert.PrintCminor.destination option_dcminor ".cm"; set_dest Vericert.PrintRTL.destination option_drtl ".rtl"; set_dest Vericert.PrintRTLBlock.destination option_drtlblock ".rtlblock"; + set_dest Vericert.PrintRTLPar.destination option_drtlpar ".rtlpar"; set_dest Vericert.PrintHTL.destination option_dhtl ".htl"; set_dest Vericert.Regalloc.destination_alloctrace option_dalloctrace ".alloctrace"; set_dest Vericert.PrintLTL.destination option_dltl ".ltl"; @@ -390,6 +391,7 @@ let cmdline_actions = Exact "-dcminor", Set option_dcminor; Exact "-drtl", Set option_drtl; Exact "-drtlblock", Set option_drtlblock; + Exact "-drtlpar", Set option_drtlpar; Exact "-dhtl", Set option_dhtl; Exact "-dltl", Set option_dltl; Exact "-dalloctrace", Set option_dalloctrace; @@ -403,6 +405,7 @@ let cmdline_actions = option_dcminor := true; option_drtl := true; option_drtlblock := true; + option_drtlpar := true; option_dhtl := true; option_dltl := true; option_dalloctrace := true; @@ -417,10 +420,10 @@ let cmdline_actions = warning_options @ (* Vericert.Interpreter mode *) [ Exact "-interp", Set option_interp; - Exact "-quiet", Unit (fun () -> Vericert.Interp.trace := 0); - Exact "-trace", Unit (fun () -> Vericert.Interp.trace := 2); - Exact "-random", Unit (fun () -> Vericert.Interp.mode := Vericert.Interp.Random); - Exact "-all", Unit (fun () -> Vericert.Interp.mode := Vericert.Interp.All) + Exact "-quiet", Unit (fun () -> Vericert.Interp.trace := 0); + Exact "-trace", Unit (fun () -> Vericert.Interp.trace := 2); + Exact "-random", Unit (fun () -> Vericert.Interp.mode := Vericert.Interp.Random); + Exact "-all", Unit (fun () -> Vericert.Interp.mode := Vericert.Interp.All) ] (* Optimization options *) (* -f options: come in -f and -fno- variants *) diff --git a/src/Compiler.v b/src/Compiler.v index d99ce56..e837e9d 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -82,6 +82,7 @@ We then need to declare the external OCaml functions used to print out intermedi Parameter print_RTL: Z -> RTL.program -> unit. Parameter print_HTL: HTL.program -> unit. Parameter print_RTLBlock: Z -> RTLBlock.program -> unit. +Parameter print_RTLPar: Z -> RTLPar.program -> unit. Definition print {A: Type} (printer: A -> unit) (prog: A) : A := let unused := printer prog in prog. diff --git a/src/VericertClflags.ml b/src/VericertClflags.ml index 534962b..930b613 100644 --- a/src/VericertClflags.ml +++ b/src/VericertClflags.ml @@ -5,5 +5,6 @@ let option_debug_hls = ref false let option_initial = ref false let option_dhtl = ref false let option_drtlblock = ref false +let option_drtlpar = ref false let option_hls_schedule = ref false let option_fif_conv = ref false diff --git a/src/extraction/Extraction.v b/src/extraction/Extraction.v index 8aec96e..5c1dac5 100644 --- a/src/extraction/Extraction.v +++ b/src/extraction/Extraction.v @@ -141,6 +141,7 @@ Extract Constant Compiler.print_Cminor => "PrintCminor.print_if". Extract Constant driver.Compiler.print_RTL => "PrintRTL.print_if". Extract Constant Compiler.print_RTL => "PrintRTL.print_if". Extract Constant Compiler.print_RTLBlock => "PrintRTLBlock.print_if". +Extract Constant Compiler.print_RTLPar => "PrintRTLPar.print_if". Extract Constant Compiler.print_HTL => "PrintHTL.print_if". Extract Constant Compiler.print_LTL => "PrintLTL.print_if". Extract Constant Compiler.print_Mach => "PrintMach.print_if". 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/Compiler.v | 4 ++ src/hls/PipelineOp.v | 141 ++++++++++++++++++++++++++++++++++++++++++++++++--- 2 files changed, 137 insertions(+), 8 deletions(-) diff --git a/src/Compiler.v b/src/Compiler.v index e837e9d..e9d76dc 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -68,6 +68,7 @@ Require vericert.hls.RTLPargen. Require vericert.hls.HTLPargen. Require vericert.hls.Pipeline. Require vericert.hls.IfConversion. +Require vericert.hls.PipelineOp. Require vericert.HLSOpts. Require Import vericert.hls.HTLgenproof. @@ -240,6 +241,9 @@ Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := @@ total_if HLSOpts.optim_if_conversion IfConversion.transf_program @@ print (print_RTLBlock 2) @@@ RTLPargen.transl_program + @@ print (print_RTLPar 1) + @@ PipelineOp.transf_program + @@ print (print_RTLPar 2) @@@ HTLPargen.transl_program @@ print print_HTL @@ Veriloggen.transl_program. 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(-) 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(-) 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(-) 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 003c2ebeed882c83ac06a0cd1f7449677ca31d14 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:00:12 +0000 Subject: Rename the verilator script --- driver/verilator_main.cpp | 35 ----------------------------------- scripts/verilator_main.cpp | 35 +++++++++++++++++++++++++++++++++++ 2 files changed, 35 insertions(+), 35 deletions(-) delete mode 100644 driver/verilator_main.cpp create mode 100644 scripts/verilator_main.cpp diff --git a/driver/verilator_main.cpp b/driver/verilator_main.cpp deleted file mode 100644 index 4561158..0000000 --- a/driver/verilator_main.cpp +++ /dev/null @@ -1,35 +0,0 @@ -#include -#include -#include "Vmain.h" -#include "verilated.h" - -int main(int argc, char **argv) { - // Initialize Verilators variables - Verilated::commandArgs(argc, argv); - - // Create an instance of our module under test - Vmain *tb = new Vmain; - - tb->clk = 0; - tb->start = 0; - tb->reset = 0; - tb->eval(); tb->clk = 1; tb->eval(); tb->clk = 0; tb->eval(); - tb->reset = 1; - tb->eval(); tb->clk = 1; tb->eval(); tb->clk = 0; tb->eval(); - tb->reset = 0; - tb->eval(); tb->clk = 1; tb->eval(); tb->clk = 0; tb->eval(); - - int cycles = 1; - - // Tick the clock until we are done - while(!tb->finish) { - tb->clk = 1; - tb->eval(); - tb->clk = 0; - tb->eval(); - cycles++; - } - - printf("cycles: %d\nfinished: %d\n", cycles, (unsigned)tb->return_val); - exit(EXIT_SUCCESS); -} diff --git a/scripts/verilator_main.cpp b/scripts/verilator_main.cpp new file mode 100644 index 0000000..4561158 --- /dev/null +++ b/scripts/verilator_main.cpp @@ -0,0 +1,35 @@ +#include +#include +#include "Vmain.h" +#include "verilated.h" + +int main(int argc, char **argv) { + // Initialize Verilators variables + Verilated::commandArgs(argc, argv); + + // Create an instance of our module under test + Vmain *tb = new Vmain; + + tb->clk = 0; + tb->start = 0; + tb->reset = 0; + tb->eval(); tb->clk = 1; tb->eval(); tb->clk = 0; tb->eval(); + tb->reset = 1; + tb->eval(); tb->clk = 1; tb->eval(); tb->clk = 0; tb->eval(); + tb->reset = 0; + tb->eval(); tb->clk = 1; tb->eval(); tb->clk = 0; tb->eval(); + + int cycles = 1; + + // Tick the clock until we are done + while(!tb->finish) { + tb->clk = 1; + tb->eval(); + tb->clk = 0; + tb->eval(); + cycles++; + } + + printf("cycles: %d\nfinished: %d\n", cycles, (unsigned)tb->return_val); + exit(EXIT_SUCCESS); +} -- cgit From b25501acffb8fa761832dfe9ad1a3138b165ff90 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:00:52 +0000 Subject: Remove verilator from Makefile --- benchmarks/polybench-syn/common.mk | 1 + scripts/verilator_main.cpp | 4 ++-- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/benchmarks/polybench-syn/common.mk b/benchmarks/polybench-syn/common.mk index 4d450e8..0913bf5 100644 --- a/benchmarks/polybench-syn/common.mk +++ b/benchmarks/polybench-syn/common.mk @@ -34,6 +34,7 @@ clean: rm -f *.clog rm -f *.tmp rm -f $(TARGETS) + rm -rf *.verilator .PRECIOUS: %.v %.gcc %.iver .PHONY: all clean diff --git a/scripts/verilator_main.cpp b/scripts/verilator_main.cpp index 4561158..94b6a33 100644 --- a/scripts/verilator_main.cpp +++ b/scripts/verilator_main.cpp @@ -19,7 +19,7 @@ int main(int argc, char **argv) { tb->reset = 0; tb->eval(); tb->clk = 1; tb->eval(); tb->clk = 0; tb->eval(); - int cycles = 1; + size_t cycles = 1; // Tick the clock until we are done while(!tb->finish) { @@ -30,6 +30,6 @@ int main(int argc, char **argv) { cycles++; } - printf("cycles: %d\nfinished: %d\n", cycles, (unsigned)tb->return_val); + printf("cycles: %lu\nfinished: %u\n", cycles, (unsigned)tb->return_val); exit(EXIT_SUCCESS); } -- cgit From 96f95a476eae57a4980f28a3cbee17ec431d5b6d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:01:03 +0000 Subject: Add if-conversion to makefile --- test/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/Makefile b/test/Makefile index 9413c70..fa482c6 100644 --- a/test/Makefile +++ b/test/Makefile @@ -1,6 +1,6 @@ CC ?= gcc VERICERT ?= vericert -VERICERT_OPTS ?= -fschedule +VERICERT_OPTS ?= -fschedule -fif-conv IVERILOG ?= iverilog IVERILOG_OPTS ?= -- 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(-) 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(-) 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(-) 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(-) 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(-) 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 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 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(-) 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 fd437f780d5b314b3e21b4125306ac9a299f08c9 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 14 Nov 2021 22:22:12 +0000 Subject: Improve Makefile for polybench --- benchmarks/polybench-syn/Makefile | 24 ++++++++++++------------ benchmarks/polybench-syn/common.mk | 4 ++-- 2 files changed, 14 insertions(+), 14 deletions(-) diff --git a/benchmarks/polybench-syn/Makefile b/benchmarks/polybench-syn/Makefile index 2c20246..7095d94 100644 --- a/benchmarks/polybench-syn/Makefile +++ b/benchmarks/polybench-syn/Makefile @@ -1,15 +1,15 @@ all: - $(MAKE) -C stencils - $(MAKE) -C medley - $(MAKE) -C linear-algebra/blas - $(MAKE) -C linear-algebra/kernels - $(MAKE) -C linear-algebra/solvers - $(MAKE) -C data-mining + -$(MAKE) -C stencils + -$(MAKE) -C medley + -$(MAKE) -C linear-algebra/blas + -$(MAKE) -C linear-algebra/kernels + -$(MAKE) -C linear-algebra/solvers + -$(MAKE) -C data-mining clean: - $(MAKE) clean -C stencils - $(MAKE) clean -C medley - $(MAKE) clean -C linear-algebra/blas - $(MAKE) clean -C linear-algebra/kernels - $(MAKE) clean -C linear-algebra/solvers - $(MAKE) clean -C data-mining + -$(MAKE) clean -C stencils + -$(MAKE) clean -C medley + -$(MAKE) clean -C linear-algebra/blas + -$(MAKE) clean -C linear-algebra/kernels + -$(MAKE) clean -C linear-algebra/solvers + -$(MAKE) clean -C data-mining diff --git a/benchmarks/polybench-syn/common.mk b/benchmarks/polybench-syn/common.mk index 0913bf5..88fb059 100644 --- a/benchmarks/polybench-syn/common.mk +++ b/benchmarks/polybench-syn/common.mk @@ -1,11 +1,11 @@ VERICERT ?= vericert -VERICERT_OPTS ?= -DSYNTHESIS -fschedule +VERICERT_OPTS ?= -DSYNTHESIS -O0 -finline IVERILOG ?= iverilog IVERILOG_OPTS ?= VERILATOR ?= verilator -VERILATOR_OPTS ?= -Wno-fatal --top main --exe /home/ymherklotz/projects/vericert/driver/verilator_main.cpp +VERILATOR_OPTS ?= -Wno-fatal --top main --exe /home/ymherklotz/projects/vericert/scripts/verilator_main.cpp TARGETS ?= -- 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(+) 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(-) 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(-) 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 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 1b0d7b8da1ab71af0577c516b2618aeb94abbc34 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 14 Nov 2021 22:23:43 +0000 Subject: Add RTLParFU to top-level --- src/Compiler.v | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/Compiler.v b/src/Compiler.v index ecea2fc..056e404 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -65,6 +65,7 @@ Require vericert.hls.HTLgen. Require vericert.hls.RTLBlock. Require vericert.hls.RTLBlockgen. Require vericert.hls.RTLPargen. +Require vericert.hls.RTLParFUgen. Require vericert.hls.HTLPargen. Require vericert.hls.Pipeline. Require vericert.hls.IfConversion. @@ -246,6 +247,7 @@ Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := @@ print (print_RTLBlock 1) @@@ RTLPargen.transl_program @@ print (print_RTLPar 0) + @@@ RTLParFUgen.transl_program @@@ HTLPargen.transl_program @@ print (print_HTL 0) @@ Veriloggen.transl_program. -- 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(-) 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(-) 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(-) 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(-) 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(-) 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(-) 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 --- driver/VericertDriver.ml | 3 +++ src/Compiler.v | 2 ++ src/VericertClflags.ml | 1 + src/extraction/Extraction.v | 1 + src/hls/HTLPargen.v | 58 +++++++++++++-------------------------------- src/hls/RTLParFUgen.v | 2 +- src/hls/Schedule.ml | 9 +++---- 7 files changed, 30 insertions(+), 46 deletions(-) diff --git a/driver/VericertDriver.ml b/driver/VericertDriver.ml index 467ae37..e89ff86 100644 --- a/driver/VericertDriver.ml +++ b/driver/VericertDriver.ml @@ -67,6 +67,7 @@ let compile_c_file sourcename ifile ofile = set_dest Vericert.PrintRTL.destination option_drtl ".rtl"; set_dest Vericert.PrintRTLBlock.destination option_drtlblock ".rtlblock"; set_dest Vericert.PrintRTLPar.destination option_drtlpar ".rtlpar"; + set_dest Vericert.PrintRTLParFU.destination option_drtlparfu ".rtlparfu"; set_dest Vericert.PrintHTL.destination option_dhtl ".htl"; set_dest Vericert.Regalloc.destination_alloctrace option_dalloctrace ".alloctrace"; set_dest Vericert.PrintLTL.destination option_dltl ".ltl"; @@ -393,6 +394,7 @@ let cmdline_actions = Exact "-drtl", Set option_drtl; Exact "-drtlblock", Set option_drtlblock; Exact "-drtlpar", Set option_drtlpar; + Exact "-drtlparfu", Set option_drtlparfu; Exact "-dhtl", Set option_dhtl; Exact "-dltl", Set option_dltl; Exact "-dalloctrace", Set option_dalloctrace; @@ -407,6 +409,7 @@ let cmdline_actions = option_drtl := true; option_drtlblock := true; option_drtlpar := true; + option_drtlparfu := true; option_dhtl := true; option_dltl := true; option_dalloctrace := true; diff --git a/src/Compiler.v b/src/Compiler.v index ff0938e..0ac2b80 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -86,6 +86,7 @@ Parameter print_RTL: Z -> RTL.program -> unit. Parameter print_HTL: Z -> HTL.program -> unit. Parameter print_RTLBlock: Z -> RTLBlock.program -> unit. Parameter print_RTLPar: Z -> RTLPar.program -> unit. +Parameter print_RTLParFU: Z -> RTLParFU.program -> unit. Definition print {A: Type} (printer: A -> unit) (prog: A) : A := let unused := printer prog in prog. @@ -249,6 +250,7 @@ Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := @@@ RTLPargen.transl_program @@ print (print_RTLPar 0) @@@ RTLParFUgen.transl_program + @@ print (print_RTLParFU 0) @@@ HTLPargen.transl_program @@ print (print_HTL 0) @@ Veriloggen.transl_program. diff --git a/src/VericertClflags.ml b/src/VericertClflags.ml index ab3c949..0402d49 100644 --- a/src/VericertClflags.ml +++ b/src/VericertClflags.ml @@ -6,6 +6,7 @@ let option_initial = ref false let option_dhtl = ref false let option_drtlblock = ref false let option_drtlpar = ref false +let option_drtlparfu = ref false let option_hls_schedule = ref false let option_fif_conv = ref false let option_fram = ref true diff --git a/src/extraction/Extraction.v b/src/extraction/Extraction.v index bca8fb5..aeefd2a 100644 --- a/src/extraction/Extraction.v +++ b/src/extraction/Extraction.v @@ -148,6 +148,7 @@ Extract Constant Compiler.print_RTLBlock => "PrintRTLBlock.print_if". Extract Constant Compiler.print_RTLPar => "PrintRTLPar.print_if". Extract Constant Compiler.print_HTL => "PrintHTL.print_if". Extract Constant Compiler.print_RTLPar => "PrintRTLPar.print_if". +Extract Constant Compiler.print_RTLParFU => "PrintRTLParFU.print_if". Extract Constant Compiler.print_LTL => "PrintLTL.print_if". Extract Constant Compiler.print_Mach => "PrintMach.print_if". Extract Constant Compiler.print => "fun (f: 'a -> unit) (x: 'a) -> f x; x". 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 82d69d247c7de8387b92936086abdc3d441c8628 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:10:55 +0000 Subject: Rename pipelining --- src/SoftwarePipelining/LICENSE | 19 - src/SoftwarePipelining/SPBase_types.ml | 128 ---- src/SoftwarePipelining/SPBasic.ml | 819 ------------------------ src/SoftwarePipelining/SPBasic.mli | 57 -- src/SoftwarePipelining/SPDebug.ml | 21 - src/SoftwarePipelining/SPIMS.ml | 189 ------ src/SoftwarePipelining/SPIMS.mli | 22 - src/SoftwarePipelining/SPMVE.ml | 299 --------- src/SoftwarePipelining/SPMVE.mli | 17 - src/SoftwarePipelining/SPSymbolic_evaluation.ml | 226 ------- src/SoftwarePipelining/SPTyping.ml | 526 --------------- src/SoftwarePipelining/SPTyping.mli | 14 - src/SoftwarePipelining/SoftwarePipelining.ml | 74 --- src/pipelining/LICENSE | 19 + src/pipelining/SPBase_types.ml | 128 ++++ src/pipelining/SPBasic.ml | 819 ++++++++++++++++++++++++ src/pipelining/SPBasic.mli | 57 ++ src/pipelining/SPDebug.ml | 21 + src/pipelining/SPIMS.ml | 189 ++++++ src/pipelining/SPIMS.mli | 22 + src/pipelining/SPMVE.ml | 299 +++++++++ src/pipelining/SPMVE.mli | 17 + src/pipelining/SPSymbolic_evaluation.ml | 226 +++++++ src/pipelining/SPTyping.ml | 526 +++++++++++++++ src/pipelining/SPTyping.mli | 14 + src/pipelining/SoftwarePipelining.ml | 74 +++ 26 files changed, 2411 insertions(+), 2411 deletions(-) delete mode 100644 src/SoftwarePipelining/LICENSE delete mode 100644 src/SoftwarePipelining/SPBase_types.ml delete mode 100644 src/SoftwarePipelining/SPBasic.ml delete mode 100644 src/SoftwarePipelining/SPBasic.mli delete mode 100644 src/SoftwarePipelining/SPDebug.ml delete mode 100644 src/SoftwarePipelining/SPIMS.ml delete mode 100644 src/SoftwarePipelining/SPIMS.mli delete mode 100644 src/SoftwarePipelining/SPMVE.ml delete mode 100644 src/SoftwarePipelining/SPMVE.mli delete mode 100644 src/SoftwarePipelining/SPSymbolic_evaluation.ml delete mode 100644 src/SoftwarePipelining/SPTyping.ml delete mode 100644 src/SoftwarePipelining/SPTyping.mli delete mode 100644 src/SoftwarePipelining/SoftwarePipelining.ml create mode 100644 src/pipelining/LICENSE create mode 100644 src/pipelining/SPBase_types.ml create mode 100644 src/pipelining/SPBasic.ml create mode 100644 src/pipelining/SPBasic.mli create mode 100644 src/pipelining/SPDebug.ml create mode 100644 src/pipelining/SPIMS.ml create mode 100644 src/pipelining/SPIMS.mli create mode 100644 src/pipelining/SPMVE.ml create mode 100644 src/pipelining/SPMVE.mli create mode 100644 src/pipelining/SPSymbolic_evaluation.ml create mode 100644 src/pipelining/SPTyping.ml create mode 100644 src/pipelining/SPTyping.mli create mode 100644 src/pipelining/SoftwarePipelining.ml diff --git a/src/SoftwarePipelining/LICENSE b/src/SoftwarePipelining/LICENSE deleted file mode 100644 index e275fa0..0000000 --- a/src/SoftwarePipelining/LICENSE +++ /dev/null @@ -1,19 +0,0 @@ -Copyright (c) 2008,2009,2010 Jean-Baptiste Tristan and INRIA - -Permission is hereby granted, free of charge, to any person obtaining a copy -of this software and associated documentation files (the "Software"), to deal -in the Software without restriction, including without limitation the rights -to use, copy, modify, merge, publish, distribute, sublicense, and/or sell -copies of the Software, and to permit persons to whom the Software is -furnished to do so, subject to the following conditions: - -The above copyright notice and this permission notice shall be included in -all copies or substantial portions of the Software. - -THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR -IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, -FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE -AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER -LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, -OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE -SOFTWARE. diff --git a/src/SoftwarePipelining/SPBase_types.ml b/src/SoftwarePipelining/SPBase_types.ml deleted file mode 100644 index ba92340..0000000 --- a/src/SoftwarePipelining/SPBase_types.ml +++ /dev/null @@ -1,128 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open Camlcoq -open Op -open Registers -open Memory -open Mem -open AST - -type ('a,'b) sum = ('a,'b) Datatypes.sum - -type instruction = - | Inop - | Iop of operation * reg list * reg - | Iload of memory_chunk * addressing * reg list * reg - | Istore of memory_chunk * addressing * reg list * reg - | Icall of signature * (reg, ident) sum * reg list * reg - | Itailcall of signature * (reg, ident) sum * reg list - | Icond of condition * reg list - | Ireturn of reg option - -type resource = Reg of reg | Mem - -let inst_coq_to_caml = function - | RTL.Inop succ -> Inop - | RTL.Iop (op, args, dst, succ) -> Iop (op, args, dst) - | RTL.Iload (chunk, mode, args, dst, succ) -> Iload (chunk, mode, args, dst) - | RTL.Istore (chunk, mode, args, src, succ) -> Istore (chunk, mode, args, src) - | RTL.Icall (sign, id, args, dst, succ) -> Icall (sign, id, args, dst) - | RTL.Itailcall (sign, id, args) -> Itailcall (sign, id, args) - | RTL.Icond (cond, args, succ1, succ2) -> Icond (cond, args) - | RTL.Ireturn (res) -> Ireturn (res) - -let inst_caml_to_coq i (links : RTL.node list) = - match i,links with - | Inop,[p] -> RTL.Inop p - | Iop (op, args, dst),[p] -> RTL.Iop (op, args, dst,p) - | Iload (chunk, mode, args, dst),[p] -> RTL.Iload (chunk, mode, args,dst,p) - | Istore (chunk, mode, args, src),[p] -> RTL.Istore (chunk, mode, args, src,p) - | Icall (sign, id, args, dst),[p] -> RTL.Icall (sign, id, args, dst,p) - | Itailcall (sign, id, args),[] -> RTL.Itailcall (sign, id, args) - | Icond (cond, args),[p1;p2] -> RTL.Icond (cond,args,p1,p2) - | Ireturn (res),[] -> RTL.Ireturn res - | _,_ -> failwith "Echec lors de la conversion des instrucitons internes vers compcert" - - -let print_inst node = string_of_int (snd node) - -let to_int = fun n -> P.to_int n -let to_binpos = fun n -> P.of_int n - -let rec string_of_args args = - match args with - | [] -> "" - | arg :: args -> "r" ^ (string_of_int (to_int arg)) ^ " " ^ string_of_args args - -let string_of_z z = string_of_int (Z.to_int z) - -let string_of_comparison = function - | Integers.Ceq -> "eq" - | Integers.Cne -> "ne" - | Integers.Clt -> "lt" - | Integers.Cle -> "le" - | Integers.Cgt -> "gt" - | Integers.Cge -> "ge" - -let string_of_cond = function - | Ccomp c -> Printf.sprintf "comp %s" (string_of_comparison c) - | Ccompu c -> Printf.sprintf "compu %s" (string_of_comparison c) - | Ccompimm (c,i) -> Printf.sprintf "compimm %s %s" (string_of_comparison c) (string_of_z i) - | Ccompuimm (c,i) -> Printf.sprintf "compuimm %s %s" (string_of_comparison c) (string_of_z i) - | Ccompf c -> Printf.sprintf "compf %s" (string_of_comparison c) - | Cnotcompf c -> Printf.sprintf "notcompf %s" (string_of_comparison c) - | Cmaskzero i -> Printf.sprintf "maskzero %s" (string_of_z i) - | Cmasknotzero i -> Printf.sprintf "masknotzero %s" (string_of_z i) - -let string_of_op = function - | Omove -> "move" - | Ointconst i -> Printf.sprintf "intconst %s" (string_of_z i) - | Ofloatconst f -> Printf.sprintf "floatconst %s" (string_of_float (camlfloat_of_coqfloat32 f)) - | Ocast8signed -> "cast8signed" - | Ocast8unsigned -> "cast8unsigned" - | Ocast16signed -> "cast16signed" - | Ocast16unsigned -> "cast16unsigned" - | Osub -> "sub" - | Omul -> "mul" - | Omulimm i -> Printf.sprintf "mulimm %s" (string_of_z i) - | Odiv -> "div" - | Odivu -> "divu" - | Oand -> "and" - | Oandimm i -> Printf.sprintf "andimm %s" (string_of_z i) - | Oor -> "or" - | Oorimm i -> Printf.sprintf "orimm %s" (string_of_z i) - | Oxor -> "xor" - | Oxorimm i -> Printf.sprintf "xorimm %s" (string_of_z i) - | Oshl -> "shl" - | Oshr -> "shr" - | Oshrimm i -> Printf.sprintf "shrimm %s" (string_of_z i) - | Oshrximm i -> Printf.sprintf "shrximm %s" (string_of_z i) - | Oshru -> "shru" - | Onegf -> "negf" - | Oabsf -> "absf" - | Oaddf -> "addf" - | Osubf -> "subf" - | Omulf -> "mulf" - | Odivf -> "divf" - | Osingleoffloat -> "singleoffloat" - | Ofloatofint -> "floatofint" - | Ocmp c -> Printf.sprintf "cmpcmpcmp" - | Olea _ -> "lea" - -let to_coq_list = function - | [] -> [] - | e :: l -> e :: l - -let to_caml_list = function - | [] -> [] - | e :: l -> e :: l diff --git a/src/SoftwarePipelining/SPBasic.ml b/src/SoftwarePipelining/SPBasic.ml deleted file mode 100644 index 32234b8..0000000 --- a/src/SoftwarePipelining/SPBasic.ml +++ /dev/null @@ -1,819 +0,0 @@ -(***********************************************************************) - (* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open RTL -open Camlcoq -open Graph.Pack.Digraph -open Op -open Registers -open Memory -open Mem -open AST -open SPBase_types -open SPSymbolic_evaluation - -type node = instruction * int -module G = Graph.Persistent.Digraph.AbstractLabeled - (struct type t = node end) - (struct type t = int let compare = compare let default = 0 end) -module Topo = Graph.Topological.Make (G) -module Dom = Graph.Dominator.Make (G) -module Index = Map.Make (struct type t = int let compare = compare end) - -let string_of_instruction node = - match G.V.label node with - | (Inop,n) -> Printf.sprintf "%i nop" n - | (Iop (op, args, dst),n) -> Printf.sprintf "%i r%i := %s %s" n (to_int dst) (string_of_op op) (string_of_args args) - | (Iload (chunk, mode, args, dst),n) -> Printf.sprintf "%i r%i := load %s" n (to_int dst) (string_of_args args) - | (Istore (chunk, mode, args, src),n) -> Printf.sprintf "%i store %i %s" n (to_int src) (string_of_args args) - | (Icall (sign, id, args, dst),n) -> Printf.sprintf "%i call" n - | (Itailcall (sign, id, args),n) -> Printf.sprintf "%i tailcall" n - (* | (Ialloc (dst, size),n) -> Printf.sprintf "%i %i := alloc" n (to_int dst) *) - | (Icond (cond, args),n) -> Printf.sprintf "%i cond %s %s" n (string_of_cond cond) (string_of_args args) - | (Ireturn (res),n) -> Printf.sprintf "%i return" n - -let string_of_node = string_of_instruction - -module Display = struct - include G - let vertex_name v = print_inst (V.label v) - let graph_attributes _ = [] - let default_vertex_attributes _ = [] - let vertex_attributes v = [`Label (string_of_instruction v)] - let default_edge_attributes _ = [] - let edge_attributes e = [`Label (string_of_int (G.E.label e))] - let get_subgraph _ = None -end -module Dot_ = Graph.Graphviz.Dot(Display) - -let dot_output g f = - let oc = open_out f in - Dot_.output_graph oc g; - close_out oc - -let display g name = - let addr = SPDebug.name ^ name in - dot_output g addr ; - ignore (Sys.command ("(dot -Tpng " ^ addr ^ " -o " ^ addr ^ ".png ; rm -f " ^ addr ^ ") & ")) - -(******************************************) - -type cfg = {graph : G.t; start : G.V.t} - -(* convert traduit un graphe RTL compcert en un graphe RTL ocamlgraph*) - -let convert f = - - let make_node inst key = - let inst = inst_coq_to_caml inst in - G.V.create (inst, to_int key) - in - - let (graph, index) = Maps.PTree.fold (fun (g,m) key inst -> - let node = make_node inst key in - (G.add_vertex g node, Index.add (to_int key) node m) - ) f.fn_code (G.empty,Index.empty) - in - - let succ = RTL.successors_map f in - let rec link n succs g = - match succs with - | [] -> g - | pos::[] -> - G.add_edge g (Index.find (to_int n) index) (Index.find (to_int pos) index) - | pos1::pos2::[] -> - let g = G.add_edge_e g (G.E.create (Index.find (to_int n) index) 1 (Index.find (to_int pos1) index)) in - G.add_edge_e g (G.E.create (Index.find (to_int n) index) 2 (Index.find (to_int pos2) index)) - | _ -> failwith "convert : trop de successeurs" - - in - - let graph = Maps.PTree.fold ( fun g key inst -> - link key (match Maps.PTree.get key succ with - Some x -> x | _ -> failwith "Could not index") g - ) f.fn_code graph - in - - {graph = graph; start = Index.find (to_int (f.fn_entrypoint)) index} - - -let convert_back g = - - G.fold_vertex (fun node m -> - let v = G.V.label node in - match (fst v) with - | Icond (_,_) -> - begin - let l = - match G.succ_e g node with - | [e1;e2] -> - if G.E.label e1 > G.E.label e2 - then [G.E.dst e2;G.E.dst e1] - else [G.E.dst e1;G.E.dst e2] - | _ -> failwith "convert_back: nombre de successeurs incoherent" - in - let succs = List.map (fun s -> to_binpos (snd (G.V.label s))) l in - Maps.PTree.set (to_binpos (snd v)) (inst_caml_to_coq (fst v) succs) m - end - | _ -> - let succs = List.map (fun s -> to_binpos (snd (G.V.label s))) (G.succ g node) in - Maps.PTree.set (to_binpos (snd v)) (inst_caml_to_coq (fst v) succs) m - ) g Maps.PTree.empty - - - - -(* dominator_tree calcule l'arbre de domination grace au code de FP *) -let dominator_tree f = - Dom.compute_idom f.graph f.start - -(* detect_loops, find the loops in the graph and returns the list of nodes in it, - in dominating order !!! This is of great importance, we suppose that it is ordered - when we build the dependency graph *) -let detect_loops graph dom_tree = - let rec is_dominating v1 v2 l = (* does v1 dominate v2 *) - match dom_tree v2 with - | v -> if v1 = v then Some (v :: l) - else is_dominating v1 v (v :: l) - | exception (Not_found) -> None - in - - G.fold_edges (fun v1 v2 loops -> - match is_dominating v2 v1 [v1] with - | None -> loops - | Some loop -> (v2,loop) :: loops - ) graph [] - -let print_index node = - Printf.printf "%i " (snd (G.V.label node)) - -let print_instruction node = - match G.V.label node with - | (Inop,n) -> Printf.printf "%i : Inop \n" n - | (Iop (op, args, dst),n) -> Printf.printf "%i : Iop \n" n - | (Iload (chunk, mode, args, dst),n) -> Printf.printf "%i : Iload \n" n - | (Istore (chunk, mode, args, src),n) -> Printf.printf "%i : Istore \n" n - | (Icall (sign, id, args, dst),n) -> Printf.printf "%i : Icall \n" n - | (Itailcall (sign, id, args),n) -> Printf.printf "%i : Itailcall \n" n - (*| (Ialloc (dst, size),n) -> Printf.printf "%i : Ialloc \n" n *) - | (Icond (cond, args),n) -> Printf.printf "%i : Icond \n" n - | (Ireturn (res),n) -> Printf.printf "%i : Ireturn \n" n - -let is_rewritten node r = - match fst (G.V.label node) with - | Inop -> false - | Iop (op, args, dst) -> if dst = r then true else false - | Iload (chunk, mode, args, dst) -> if dst = r then failwith "J'ai degote une boucle ZARBI !!!" else false - | Istore (chunk, mode, args, src) -> false - | Icall (sign, id, args, dst) -> failwith "call in a loop" - | Itailcall (sign, id, args) -> failwith "tailcall in a loop" - (* | Ialloc (dst, size) -> if dst = r then failwith "J'ai degote une boucle ZARBI !!!" else false *) - | Icond (cond, args) -> false - | Ireturn (res) -> failwith "return in a loop" - -let is_variant r loop = - List.fold_right (fun node b -> - is_rewritten node r || b - ) loop false - - -let is_pipelinable loop = (* true if loop is innermost and without control *) - - let is_acceptable node = - match fst (G.V.label node) with - | Icall _ | Itailcall _ | Ireturn _ | Icond _ (*| Ialloc _*) | Iop ((Ocmp _),_,_)-> false - | _ -> true - in - - let is_branching node = - match fst (G.V.label node) with - | Icond _ -> true - | _ -> false - in - - let is_nop node = - match fst (G.V.label node) with - | Inop -> true - | _ -> false - in - - let is_OK_aux l = - List.fold_right (fun n b -> is_acceptable n && b) l true - in - - let is_bounded node loop = - match G.V.label node with - | (Icond (cond, args),n) -> - let args = to_caml_list args in - begin - match args with - | [] -> false - | r :: [] -> is_variant r loop (* used to be not *) - | r1 :: r2 :: [] -> - begin - match is_variant r1 loop, is_variant r2 loop with - | true, true -> false - | false, true -> true - | true, false -> true - | false, false -> false - end - | _ -> false - end - | _ -> false - in - - match List.rev loop with - | v2 :: v1 :: l -> ((*Printf.printf "is_nop: %s | " (is_nop v1 |> string_of_bool);*) - Printf.printf "is_branching: %s | " (is_branching v2 |> string_of_bool); - Printf.printf "is_OK_aux: %s | " (is_OK_aux l |> string_of_bool); - Printf.printf "is_bounded: %s\n" (is_bounded v2 loop |> string_of_bool); - (*is_nop v1 && *)is_branching v2 && is_OK_aux l && is_bounded v2 loop) - | _ -> false - -let print_loops loops = - List.iter (fun loop -> print_index (fst(loop)); - print_newline (); - List.iter print_index (snd(loop)); - print_newline (); - if is_pipelinable (snd(loop)) then print_string "PIPELINABLE !" else print_string "WASTE"; - print_newline (); - List.iter print_instruction (snd(loop)); - print_newline () - ) loops - -(* type resource = Reg of reg | Mem *) -module Sim = Map.Make (struct type t = resource let compare = compare end) - -let map_get key map = - try Some (Sim.find key map) - with - | Not_found -> None - -let rec to_res l = - match l with - | [] -> [] - | e :: l -> Reg e :: to_res l - -let resources_reads_of node = - match fst (G.V.label node) with - | Inop -> [] - | Iop (op, args, dst) -> to_res args - | Iload (chunk, mode, args, dst) -> Mem :: (to_res args) - | Istore (chunk, mode, args, src) -> Mem :: Reg src :: (to_res args) - | Icall (sign, id, args, dst) -> failwith "Resource read of call" - | Itailcall (sign, id, args) -> failwith "Resource read of tailcall" - (*| Ialloc (dst, size) -> [Mem] *) - | Icond (cond, args) -> to_res args - | Ireturn (res) -> failwith "Resource read of return" - -let resources_writes_of node = - match fst (G.V.label node) with - | Inop -> [] - | Iop (op, args, dst) -> [Reg dst] - | Iload (chunk, mode, args, dst) -> [Reg dst] - | Istore (chunk, mode, args, src) -> [Mem] - | Icall (sign, id, args, dst) -> failwith "Resource read of call" - | Itailcall (sign, id, args) -> failwith "Resource read of tailcall" - (*| Ialloc (dst, size) -> (Reg dst) :: [Mem]*) - | Icond (cond, args) -> [] - | Ireturn (res) -> failwith "Resource read of return" - -let build_intra_dependency_graph loop = - - let rec build_aux graph read write l = - match l with - | [] -> (graph,(read,write)) - | v :: l-> - let g = G.add_vertex graph v in - let reads = resources_reads_of v in - let writes = resources_writes_of v in - (* dependances RAW *) - let g = List.fold_right (fun r g -> - match map_get r write with - | Some n -> G.add_edge_e g (G.E.create n 1 v) - | None -> g - ) reads g in - (* dependances WAR *) - let g = List.fold_right (fun r g -> - match map_get r read with - | Some l -> List.fold_right (fun n g -> G.add_edge_e g (G.E.create n 2 v)) l g - | None -> g - ) writes g in - (* dependances WAW *) - let g = List.fold_right (fun r g -> - match map_get r write with - | Some n -> G.add_edge_e g (G.E.create n 3 v) - | None -> g - ) writes g in - let write = List.fold_right (fun r m -> Sim.add r v m) writes write in - let read_tmp = List.fold_right (fun r m -> - match map_get r read with - | Some al -> Sim.add r (v :: al) m - | None -> Sim.add r (v :: []) m - ) reads read - in - let read = List.fold_right (fun r m -> Sim.add r [] m) writes read_tmp in - - build_aux g read write l - in - - build_aux G.empty Sim.empty Sim.empty (List.tl loop) - -let build_inter_dependency_graph loop = - - let rec build_aux2 graph read write l = - match l with - | [] -> graph - | v :: l-> - let g = graph in - let reads = resources_reads_of v in - let writes = resources_writes_of v in - (* dependances RAW *) - let g = List.fold_right (fun r g -> - match map_get r write with - | Some n -> (* if n = v then g else *) G.add_edge_e g (G.E.create n 4 v) - | None -> g - ) reads g in - (* dependances WAR *) - let g = List.fold_right (fun r g -> - match map_get r read with - | Some l -> List.fold_right - (fun n g -> (* if n = v then g else *) G.add_edge_e g (G.E.create n 5 v)) l g - | None -> g - ) writes g in - (* dependances WAW *) - let g = List.fold_right (fun r g -> - match map_get r write with - | Some n -> (* if n = v then g else *) G.add_edge_e g (G.E.create n 6 v) - | None -> g - ) writes g in - let write = List.fold_right (fun r m -> Sim.remove r m) writes write in - let read = List.fold_right (fun r m -> Sim.remove r m) writes read in - - - build_aux2 g read write l - in - - let (g,(r,w)) = build_intra_dependency_graph loop in - build_aux2 g r w (List.tl loop) - -(* patch_graph prepare le graphe pour la boucle commencant au noeud entry - qui a une borne de boucle bound et pour un software pipelining - de au minimum min tours et de deroulement ur *) -(* this is rather technical so we give many comments *) - -(* let n1 = G.V.create (Iop ((Ointconst ur),to_coq_list [],r1),next_pc) in *) -(* let next_pc = next_pc + 1 in *) -(* let n2 = G.V.create (Iop (Odiv,to_coq_list [bound;r1],r2),next_pc) in *) -(* let next_pc = next_pc + 1 in *) -(* let n3 = G.V.create (Iop ((Omulimm ur),to_coq_list [r2],r3),next_pc) in *) -(* let next_pc = next_pc + 1 in *) -(* let n4 = G.V.create (Iop (Osub,to_coq_list [bound;r3],r4),next_pc) in *) -(* let next_pc = next_pc + 1 in *) -(* let n5 = G.V.create (Iop (Omove,to_coq_list [r3],bound),next_pc) in (\* retouchee, [r3],bound *\) *) - - -let patch_graph graph entry lastone loop bound min ur r1 r2 r3 r4 next_pc prolog epilog ramp_up ramp_down = - - (* 1. Break the edge that enters the loop, except for the backedge *) - let preds_e = G.pred_e graph entry in - let wannabes = List.map G.E.src preds_e in - let wannabes = List.filter (fun e -> not (e = lastone)) wannabes in - let graph = List.fold_right (fun e g -> G.remove_edge_e g e) preds_e graph in - let graph = G.add_edge graph lastone entry in - - (* 2. Add the test for minimal iterations and link it*) - - let cond = G.V.create (Icond ((Ccompimm (Integers.Cle,min)),to_coq_list [bound]), next_pc) in - let graph = G.add_vertex graph cond in - let next_pc = next_pc + 1 in - - (* 3. Link its predecessors and successors *) - (* It is false in case there is a condition that points to the entry: - inthis case, the edge should not be labeled with 0 !*) - - let graph = List.fold_right (fun n g -> G.add_edge g n cond) wannabes graph in - let graph = G.add_edge_e graph (G.E.create cond 1 entry) in - - - (* 4. Add the div and modulo code, link it *) - let n1 = G.V.create (Iop ((Ointconst ur),to_coq_list [],r1),next_pc) in - let next_pc = next_pc + 1 in - let n2 = G.V.create (Iop (Odiv,to_coq_list [bound;r1],r2),next_pc) in - let next_pc = next_pc + 1 in - let n3 = G.V.create (Iop (Omove,to_coq_list [bound],r4),next_pc) in - let next_pc = next_pc + 1 in - let n4 = G.V.create (Iop ((Omulimm ur ),to_coq_list [r2],r3),next_pc) in - let next_pc = next_pc + 1 in - let n5 = G.V.create (Iop ((Olea (Aindexed (Z.of_sint (-1)))),to_coq_list [r3],bound),next_pc) in (* retouchee, [r3],bound *) - let next_pc = next_pc + 1 in - let graph = G.add_vertex graph n1 in - let graph = G.add_vertex graph n2 in - let graph = G.add_vertex graph n3 in - let graph = G.add_vertex graph n4 in - let graph = G.add_vertex graph n5 in - let graph = G.add_edge_e graph (G.E.create cond 2 n1) in - let graph = G.add_edge graph n1 n2 in - let graph = G.add_edge graph n2 n3 in - let graph = G.add_edge graph n3 n4 in - let graph = G.add_edge graph n4 n5 in - - (* 5. Fabriquer la pipelined loop et la linker, sans la condition d entree *) - - let (graph,next_pc,l) = List.fold_right (fun e (g,npc,l) -> - let n = G.V.create (e,npc) in - (G.add_vertex g n, npc+1, n :: l) - ) loop (graph,next_pc,[]) in - - let pipe_cond = List.hd l in - let pipeline = List.tl l in - - let rec link l graph node = - match l with - | n1 :: n2 :: l -> link (n2 :: l) (G.add_edge graph n1 n2) node - | n1 :: [] -> G.add_edge graph n1 node - | _ -> graph - in - - let graph = link pipeline graph pipe_cond in - - (* link de l entree de la boucle *) - - let (graph,next_pc,prolog) = List.fold_right (fun e (g,npc,l) -> - let n = G.V.create (e,npc) in - (G.add_vertex g n, npc+1, n :: l) - ) prolog (graph,next_pc,[]) in - - let (graph,next_pc,epilog) = List.fold_right (fun e (g,npc,l) -> - let n = G.V.create (e,npc) in - (G.add_vertex g n, npc+1, n :: l) - ) epilog (graph,next_pc,[]) in - - (* 6. Creation du reste et branchement et la condition *) - let n6 = G.V.create (Iop (Omove,to_coq_list [r4],bound),next_pc) in (* Iop (Omove,to_coq_list [r4],bound) *) - let next_pc = next_pc + 1 in - - (* 7. Creation du ramp up *) - let ramp_up = List.map (fun (a,b) -> Iop (Omove, [b], a)) ramp_up in - let (graph,next_pc,ramp_up) = List.fold_right (fun e (g,npc,l) -> - let n = G.V.create (e,npc) in - (G.add_vertex g n, npc+1, n :: l) - ) ramp_up (graph,next_pc,[]) in - - let next_pc = next_pc + 1 in - - let ramp_down = List.map (fun (a,b) -> Iop (Omove,[b],a)) ramp_down in - let (graph,next_pc,ramp_down) = List.fold_right (fun e (g,npc,l) -> - let n = G.V.create (e,npc) in - (G.add_vertex g n, npc+1, n :: l) - ) ramp_down (graph,next_pc,[]) in - - (* let next_pc = next_pc + 1 in *) - - (* Creation des proloque et epilogue *) - - let graph = link prolog graph pipe_cond in - let graph = link ramp_up graph (List.hd prolog) in - let graph = link epilog graph (List.hd ramp_down) in - let graph = link ramp_down graph n6 in - - let graph = G.add_edge graph n5 (List.hd ramp_up) in - let graph = G.add_edge_e graph (G.E.create pipe_cond 1 (List.hd epilog)) in - let graph = G.add_edge_e graph (G.E.create pipe_cond 2 (List.hd pipeline)) in - - (* 8. Retour sur la boucle classique *) - let graph = G.add_edge graph n6 entry in - - graph - -let regs_of_node node = - match G.V.label node with - | (Inop,n) -> [] - | (Iop (op, args, dst),n) -> dst :: (to_caml_list args) - | (Iload (chunk, mode, args, dst),n) -> dst :: (to_caml_list args) - | (Istore (chunk, mode, args, src),n) -> src :: (to_caml_list args) - | (Icall (sign, id, args, dst),n) -> dst :: (to_caml_list args) - | (Itailcall (sign, id, args),n) -> (to_caml_list args) - (*| (Ialloc (dst, size),n) -> [dst]*) - | (Icond (cond, args),n) -> (to_caml_list args) - | (Ireturn (res),n) -> match res with Some res -> [res] | _ -> [] - -let max_reg_of_graph graph params = - Printf.fprintf SPDebug.dc "Calcul du registre de depart.\n"; - let regs = G.fold_vertex (fun node regs -> - (regs_of_node node) @ regs - ) graph [] in - let regs = regs @ params in - let max_reg = List.fold_right (fun reg max -> - Printf.fprintf SPDebug.dc "%i " (P.to_int reg); - if Int32.compare (P.to_int32 reg) max > 0 - then (P.to_int32 reg) - else max - ) regs Int32.zero in - - Printf.fprintf SPDebug.dc "MAX REG = %i\n" (Int32.to_int max_reg); - BinPos.Pos.succ (P.of_int32 max_reg) - -let get_bound node loop = - match G.V.label node with - | (Icond (cond, args),n) -> - let args = to_caml_list args in - begin - match args with - | [] -> failwith "get_bound: condition sans variables" - | r :: [] -> if is_variant r loop then failwith "Pas de borne dans la boucle" else r (* Modified false to true condition. *) - | r1 :: r2 :: [] -> - begin - match is_variant r1 loop, is_variant r2 loop with - | true, true -> failwith "Pas de borne dans la boucle " - | false, true -> r1 - | true, false -> r2 - | false, false -> failwith "deux bornes possibles dans la boucle" - end - | _ -> failwith "get_bound: condition avec nombre de variables superieur a 2" - end - | _ -> failwith "get_bound: the node I was given is not a condition\n" - -let get_nextpc graph = - (G.fold_vertex (fun node max -> - if (snd (G.V.label node)) > max - then (snd (G.V.label node)) - else max - ) graph 0) + 1 - -let substitute_pipeline graph loop steady_state prolog epilog min unrolling ru rd params = - let n1 = max_reg_of_graph graph params in - let n2 = (BinPos.Pos.succ n1) in - let n3 = (BinPos.Pos.succ n2) in - let n4 = (BinPos.Pos.succ n3) in - let way_in = (List.hd loop) in - let way_out = (List.hd (List.rev loop)) in - let bound = (get_bound way_out loop) in - let min = Z.of_sint min in - let unrolling = Z.of_sint unrolling in - let next_pc = get_nextpc graph in - patch_graph graph way_in way_out steady_state bound min unrolling n1 n2 n3 n4 next_pc prolog epilog ru rd - -let get_loops cfg = - let domi = dominator_tree cfg in - let loops = detect_loops cfg.graph domi in - print_loops loops; - let loops = List.filter (fun loop -> is_pipelinable (snd (loop))) loops in - loops - -type pipeline = {steady_state : G.V.t list; prolog : G.V.t list; epilog : G.V.t list; - min : int; unrolling : int; ramp_up : (reg * reg) list; ramp_down : (reg * reg) list} - -let delete_indexes l = List.map (fun e -> fst (G.V.label e) ) l - -type reg = Registers.reg - -let fresh = ref BinNums.Coq_xH - -let distance e = - match G.E.label e with - | 1 | 2 | 3 -> 0 - | _ -> 1 - -type et = IntraRAW | IntraWAW | IntraWAR | InterRAW | InterWAW | InterWAR - -let edge_type e = - match G.E.label e with - | 1 -> IntraRAW - | 2 -> IntraWAR - | 3 -> IntraWAW - | 4 -> InterRAW - | 5 -> InterWAR - | 6 -> InterWAW - | _ -> failwith "Unknown edge type" - -let latency n = (* A raffiner *) - match fst (G.V.label n) with - | Iop (op,args, dst) -> - begin - match op with - | Omove -> 1 - (*| Oaddimm _ -> 1*) - (*| Oadd -> 2*) - | Omul -> 4 - | Odiv -> 30 - | Omulimm _ -> 4 - | _ -> 2 - end - | Iload _ -> 1 - (* | Ialloc _ -> 20*) - | _ -> 1 - -let reforge_writes inst r = - G.V.create ((match fst (G.V.label inst) with - | Inop -> Inop - | Iop (op, args, dst) -> Iop (op, args, r) - | Iload (chunk, mode, args, dst) -> Iload (chunk, mode, args, r) - | Istore (chunk, mode, args, src) -> Istore (chunk, mode, args, src) - | Icall (sign, id, args, dst) -> failwith "reforge_writes: call" - | Itailcall (sign, id, args) -> failwith "reforge_writes: tailcall" - (* | Ialloc (dst, size) -> Ialloc (r, size)*) - | Icond (cond, args) -> Icond (cond, args) - | Ireturn (res) -> failwith "reforge_writes: return") - , snd (G.V.label inst)) - -let rec reforge_args args oldr newr = - match args with - | [] -> [] - | e :: l -> (if e = oldr then newr else e) :: (reforge_args l oldr newr) - -let rec mem_args args r = - match args with - | [] -> false - | e :: l -> if e = r then true else mem_args l r - -let check_read_exists inst r = - match fst (G.V.label inst) with - | Inop -> false - | Iop (op, args, dst) -> mem_args args r - | Iload (chunk, mode, args, dst) -> mem_args args r - | Istore (chunk, mode, args, src) -> src = r || mem_args args r - | Icall (sign, id, args, dst) -> mem_args args r - | Itailcall (sign, id, args) -> false - (*| Ialloc (dst, size) -> false*) - | Icond (cond, args) -> mem_args args r - | Ireturn (res) -> false - -let reforge_reads inst oldr newr = - assert (check_read_exists inst oldr); - G.V.create ((match fst (G.V.label inst) with - | Inop -> Inop - | Iop (op, args, dst) -> Iop (op, reforge_args args oldr newr, dst) - | Iload (chunk, mode, args, dst) -> Iload (chunk, mode, reforge_args args oldr newr, dst) - | Istore (chunk, mode, args, src) -> Istore (chunk, mode, reforge_args args oldr newr , if src = oldr then newr else src) - | Icall (sign, id, args, dst) -> failwith "reforge_reads: call" - | Itailcall (sign, id, args) -> failwith "reforge_reads: tailcall" - (*| Ialloc (dst, size) -> Ialloc (dst, size)*) - | Icond (cond, args) -> Icond (cond, reforge_args args oldr newr) - | Ireturn (res) -> failwith "reforge_reads: return") - , snd (G.V.label inst)) - -let get_succs_raw ddg node = - let succs = G.succ_e ddg node in - let succs = List.filter (fun succ -> - match G.E.label succ with - | 1 | 4 -> true - | _ -> false - ) succs in - List.map (fun e -> G.E.dst e) succs - -let written inst = - match fst (G.V.label inst) with - | Inop -> None - | Iop (op, args, dst) -> Some dst - | Iload (chunk, mode, args, dst) -> Some dst - | Istore (chunk, mode, args, src) -> None - | Icall (sign, id, args, dst) -> failwith "written: call" - | Itailcall (sign, id, args) -> failwith "written: tailcall" - (*| Ialloc (dst, size) -> Some dst*) - | Icond (cond, args) -> None - | Ireturn (res) -> failwith "written: return" - -let fresh_regs n = - let t = Array.make n (BinNums.Coq_xH) in - for i = 0 to (n - 1) do - Array.set t i (!fresh); - fresh := BinPos.Pos.succ !fresh - done; - t - -let print_reg r = Printf.fprintf SPDebug.dc "%i " (P.to_int r) - -let is_cond node = - match fst (G.V.label node) with - | Icond _ -> true - | _ -> false - - -(*******************************************) - -let watch_regs l = List.fold_right (fun (a,b) l -> - if List.mem a l then l else a :: l - ) l [] - -let make_moves = List.map (fun (a,b) -> Iop (Omove,[b],a)) - -let rec repeat l n = - match n with - | 0 -> [] - | n -> l @ repeat l (n-1) - -let fv = ref 0 - -let apply_pipeliner f p ?(debug=false) = - Printf.fprintf SPDebug.dc "******************** NEW FUNCTION ***********************\n"; - let cfg = convert f in - incr fv; - if debug then display cfg.graph ("input" ^ (string_of_int !fv)); - let loops = get_loops cfg in - Printf.fprintf SPDebug.dc "Loops: %d\n" (List.length loops); - let ddgs = List.map (fun (qqch,loop) -> (loop,build_inter_dependency_graph loop)) loops in - - let lv = ref 0 in - - let graph = List.fold_right (fun (loop,ddg) graph -> - Printf.fprintf SPDebug.dc "__________________ NEW LOOP ____________________\n"; - Printf.printf "Pipelinable loop: "; - incr lv; - fresh := (BinPos.Pos.succ - (BinPos.Pos.succ - (BinPos.Pos.succ - (BinPos.Pos.succ - (BinPos.Pos.succ - (max_reg_of_graph graph (to_caml_list f.fn_params) - )))))); - Printf.fprintf SPDebug.dc "FRESH = %i \n" - (P.to_int !fresh); - match p ddg with - | Some pipe -> - Printf.printf "Rock On ! Min = %i - Unroll = %i\n" pipe.min pipe.unrolling; - let p = (make_moves pipe.ramp_up) @ (delete_indexes pipe.prolog) in - let e = (delete_indexes pipe.epilog) @ (make_moves pipe.ramp_down) in - let b = delete_indexes (List.tl (List.rev loop)) in - let bt = (List.tl (delete_indexes pipe.steady_state)) in - let cond1 = fst (G.V.label (List.hd (List.rev loop))) in - let cond2 = (List.hd (delete_indexes pipe.steady_state)) in - - let bu = symbolic_evaluation (repeat b (pipe.min + 1)) in - let pe = symbolic_evaluation (p @ e) in - let bte = symbolic_evaluation (bt @ e) in - let ebu = symbolic_evaluation (e @ repeat b pipe.unrolling) in - let regs = watch_regs pipe.ramp_down in - let c1 = symbolic_condition cond1 (repeat b pipe.unrolling) in - let d1 = symbolic_condition cond1 (repeat b (pipe.min + 1)) in - (*let c2 = symbolic_condition cond2 p in - let d2 = symbolic_condition cond2 ((make_moves pipe.ramp_up) @ bt) in*) - - - - let sbt = symbolic_evaluation (bt) in - let sep = symbolic_evaluation (e @ repeat b (pipe.unrolling - (pipe.min + 1)) @ p) in (* er @ pr *) - - Printf.printf "Initialisation : %s \n" - (if symbolic_equivalence bu pe regs then "OK" else "FAIL"); - Printf.printf "Etat stable : %s \n" - (if symbolic_equivalence bte ebu regs then "OK" else "FAIL"); - Printf.printf "Egalite fondamentale : %s \n" - (if symbolic_equivalence sbt sep (watch_regs pipe.ramp_up) then "OK" else "FAIL"); - (* Printf.printf "Condition initiale : %s \n" - (if c1 = c2 then "OK" else "FAIL"); - Printf.printf "Condition stable : %s \n" - (if d1 = d2 then "OK" else "FAIL"); - - - Printf.fprintf SPDebug.dc "pbte\n";*) - List.iter (fun e -> - Printf.fprintf SPDebug.dc "%s\n" - (string_of_node (G.V.create (e,0))) - ) (p @ bt @ e); - Printf.fprintf SPDebug.dc "bu\n"; - List.iter (fun e -> Printf.fprintf SPDebug.dc "%s\n" - (string_of_node (G.V.create (e,0))) - ) (repeat b (pipe.unrolling + pipe.min)); - - - - if debug then - display_st ("pbte"^ (string_of_int !fv) ^ (string_of_int !lv)) (p @ bt @ e) (watch_regs pipe.ramp_down); - if debug then - display_st ("bu"^ (string_of_int !fv) ^ (string_of_int !lv)) (repeat b (pipe.min + pipe.unrolling)) (watch_regs pipe.ramp_down); - - if debug then display_st ("bt"^ (string_of_int !fv) ^ (string_of_int !lv)) (bt) (watch_regs pipe.ramp_up); - if debug then display_st ("ep"^ (string_of_int !fv) ^ (string_of_int !lv)) (e @ repeat b (pipe.unrolling - (pipe.min + 1)) @ p) (watch_regs pipe.ramp_up); - - substitute_pipeline graph loop - (delete_indexes pipe.steady_state) (delete_indexes pipe.prolog) - (delete_indexes pipe.epilog) (pipe.min + pipe.unrolling) - pipe.unrolling pipe.ramp_up - pipe.ramp_down - (to_caml_list f.fn_params) - | None -> Printf.printf "Damn It ! \n"; graph - ) ddgs cfg.graph in - - if debug then display graph ("output"^ (string_of_int !fv)); - let tg = convert_back graph in - - let tg_to_type = {fn_sig = f.fn_sig; - fn_params = f.fn_params; - fn_stacksize = f.fn_stacksize; - fn_code = tg; - fn_entrypoint = f.fn_entrypoint; - (*fn_nextpc = P.of_int ((get_nextpc (graph)))*) - } in - (*SPTyping.type_function tg_to_type;*) - - tg_to_type diff --git a/src/SoftwarePipelining/SPBasic.mli b/src/SoftwarePipelining/SPBasic.mli deleted file mode 100644 index f16e524..0000000 --- a/src/SoftwarePipelining/SPBasic.mli +++ /dev/null @@ -1,57 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open Graph.Pack.Digraph - -(* DATA DEPENDENCY GRAPHS *) -module G : Graph.Sig.P - -(* We abstract the register type to make sure that the user won't mess up *) -type reg - -(* The scheduling should return a value of type pipeline *) -type pipeline = { - steady_state : G.V.t list; - prolog : G.V.t list; - epilog : G.V.t list; - min : int; - unrolling : int; - ramp_up : (reg * reg) list; - ramp_down : (reg * reg) list -} - -(* Operations on ddg *) - -val display : G.t -> string -> unit -val apply_pipeliner : RTL.coq_function -> (G.t -> pipeline option) -> ?debug:bool -> RTL.coq_function -val get_succs_raw : G.t -> G.V.t -> G.V.t list - -(* Operations on instructions, the nodes of the data dependency graph *) - -val string_of_node : G.V.t -> string -val latency : G.V.t -> int -val reforge_reads : G.V.t -> reg -> reg -> G.V.t -val reforge_writes : G.V.t -> reg -> G.V.t -val written : G.V.t -> reg option -val is_cond : G.V.t -> bool - -(* Operations on dependencies, the edges of the data dependency graph *) - -type et = IntraRAW | IntraWAW | IntraWAR | InterRAW | InterWAW | InterWAR - -val edge_type : G.E.t -> et -val distance : G.E.t -> int - -(* Getting fresh registers, int is the number of required registers *) - -val fresh_regs : int -> reg array -val print_reg : reg -> unit diff --git a/src/SoftwarePipelining/SPDebug.ml b/src/SoftwarePipelining/SPDebug.ml deleted file mode 100644 index 34d1c0c..0000000 --- a/src/SoftwarePipelining/SPDebug.ml +++ /dev/null @@ -1,21 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open Unix - -let tm = localtime (time ());; -let name = "debug/" ^ (string_of_int tm.tm_year) ^ "-" ^(string_of_int tm.tm_mon) ^ "-" ^(string_of_int tm.tm_mday) ^ - "-" ^(string_of_int tm.tm_hour) ^"-" ^(string_of_int tm.tm_min) ^ "-" ^(string_of_int tm.tm_sec) ^ "/";; -mkdir name 0o777;; -Printf.printf "Debug informations: %s \n" name ;; -let dc = open_out (name ^ "debug.log");; -let () = at_exit(fun () -> close_out dc);; diff --git a/src/SoftwarePipelining/SPIMS.ml b/src/SoftwarePipelining/SPIMS.ml deleted file mode 100644 index 0e19dec..0000000 --- a/src/SoftwarePipelining/SPIMS.ml +++ /dev/null @@ -1,189 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open Graph.Pack.Digraph -open SPBasic - -module NI = Map.Make (struct type t = G.V.t let compare = compare end) - -let find key map def = - try NI.find key map - with - | Not_found -> def - -let unpack v = - match v with - | Some v -> v - | None -> failwith "unpack abusif" - -let dep_latency edge = - match edge_type edge with - | IntraRAW | InterRAW -> latency (G.E.src edge) - | _ -> 1 - -let estart ddg schedule node ii = - let preds = G.pred_e ddg node in - let starts = List.map (fun edge -> - match find (G.E.src edge) schedule None with - | Some t -> - let start = t + dep_latency edge - ii * distance edge in - (*Printf.printf "start : %i \n" start;*) - if start < 0 then 0 else start - | None -> 0 - ) preds in - List.fold_left (fun max e -> if max > e then max else e) 0 starts - -let resource_conflict time mrt ii = - match Array.get mrt (time mod ii) with - | None -> false - | Some v -> true - -let rec scan_time time maxtime mrt ii = - if time <= maxtime - then - begin - if resource_conflict time mrt ii - then scan_time (time + 1) maxtime mrt ii - else Some time - end - else None - -let finished ddg schedule = - let unscheduled = G.fold_vertex (fun node l -> - match find node schedule None with - | Some v -> l - | None -> node :: l - ) ddg [] in - (* Printf.printf "R %i R \n" (List.length unscheduled); *) - if List.length unscheduled = 0 then true else false - -let bad_successors ddg schedule node ii = - let succs = G.succ_e ddg node in (* Le succs_ddg initial *) - (* let reftime = NI.find node schedule in *) - (* let succs_sched = NI.fold (fun node time succs -> *) - (* if time > reftime then node :: succs else succs *) - (* ) schedule [] in *) - (* let succs = List.filter (fun edge -> List.mem (G.E.dst edge) succs_sched) succs_ddg in*) - List.fold_right (fun edge bad -> - match find (G.E.dst edge) schedule None with - | Some t -> - if unpack (NI.find node schedule) + dep_latency edge - ii * distance edge > t - then (G.E.dst edge) :: bad - else bad - | None -> bad - ) succs [] - -let get_condition ddg = - let cond = G.fold_vertex (fun node cond -> - if is_cond node then Some node else cond - ) ddg None in - match cond with - | Some cond -> cond - | None -> failwith "The loop does not contain a condition. Aborting\n" - -(* Perform iterative modulo scheduling, using a heuristic for the next instruction to schedule - * [heightr], the data dependency graph to schedule [ddg], the minimum II [min_ii] and the maximum - * II [max_interval]. - *) -let modulo_schedule heightr ddg min_ii max_interval = - - let ii = ref (min_ii - 1) in - let notfound = ref true in - let sched = ref NI.empty in - - let cond = get_condition ddg in - - while (!ii < max_interval && !notfound) do - (* Printf.printf "."; flush stdout; *) - ii := !ii + 1; - (* Printf.printf "NOUVEAU II %i \n" !ii; *) - let budget = ref (G.nb_vertex ddg * 10) in - let lasttime = ref NI.empty in - (* Create the map with schedules, and add the schedule for the condition. This should go at the - * end, but in this case is set to the start. *) - let schedtime = ref (NI.add cond (Some 0) NI.empty) in - (* Create an array that is as large as the current II, which will determine where each - * instruction will be placed. *) - let mrt = Array.make !ii None in - (* Set the condition to be the initial instruction at time 0. *) - Array.set mrt 0 (Some cond); - - while (!budget > 0 && not (finished ddg !schedtime)) do (* Pretty inefficient *) - budget := !budget - 1; - (* Get next instruction to schedule. *) - let h = heightr ddg !schedtime in - let mintime = estart ddg !schedtime h !ii in - (* Printf.printf "tmin (%s) = %i \n" (string_of_node h) mintime; *) - let maxtime = mintime + !ii -1 in - let time = - match scan_time mintime maxtime mrt !ii with - | Some t -> t - | None -> (*Printf.printf "backtrack" ; *) - if mintime = 0 then 1 + find h !lasttime 0 - else max mintime (1 + find h !lasttime 0) - in - (* Printf.printf "Chosen time for %s : %i \n" (string_of_node h) time; *) - schedtime := NI.add h (Some time) !schedtime; - lasttime := NI.add h time !lasttime; - - let killed = bad_successors ddg !schedtime h !ii in - List.iter (fun n -> (* Printf.printf "Killing %s" (string_of_node n) ; *)schedtime := NI.add n None !schedtime) killed; - - begin - match Array.get mrt (time mod !ii) with - | None -> Array.set mrt (time mod !ii) (Some h) - | Some n -> - begin - (*Printf.printf "Deleting : %s \n" (string_of_node n); *) - (* Printf.printf "."; *) - schedtime := NI.add n None !schedtime; - Array.set mrt (time mod !ii) (Some h) - end - end; - (* if finished ddg !schedtime then Printf.printf "Fini ! \n" *) - - done; - - let success = G.fold_vertex (fun node b -> - b && - match find node !schedtime None with - | Some _ -> true - | None -> false - ) ddg true in - - if success then (notfound := false; sched := !schedtime); - - done; - - if (not !notfound) - then (!sched,!ii) - else failwith "IMS failure" - -(* Take the number of vertices as the minimum resource-constrained II. However, the II might - actually be less than that in some cases, so this should be adjusted accordingly. *) -let res_m_ii ddg = - G.nb_vertex ddg - -let pipeliner ddg heightr = - let mii = res_m_ii ddg in - Printf.fprintf SPDebug.dc "MII: %i\n" mii; - let (schedule,ii) = modulo_schedule heightr ddg mii (10 * mii) in - (NI.fold (fun n v map -> - match v with - | Some v -> NI.add n v map - | None -> failwith "pipeliner: schedule unfinished" - ) schedule NI.empty,ii) - -let print_schedule sched = - NI.iter (fun node time -> - Printf.fprintf SPDebug.dc "%s |---> %i \n" (string_of_node node) time - ) sched diff --git a/src/SoftwarePipelining/SPIMS.mli b/src/SoftwarePipelining/SPIMS.mli deleted file mode 100644 index 7c1d9a7..0000000 --- a/src/SoftwarePipelining/SPIMS.mli +++ /dev/null @@ -1,22 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open Graph.Pack.Digraph -open SPBasic - -module NI : Map.S with type key = SPBasic.G.V.t - -(* piepeliner takes a data dependency graph and returns a schedule with an initiation interval - fails if cannot find any schedule *) -val pipeliner : G.t -> (G.t -> int option NI.t -> G.V.t) -> int NI.t * int - -val print_schedule : int NI.t -> unit diff --git a/src/SoftwarePipelining/SPMVE.ml b/src/SoftwarePipelining/SPMVE.ml deleted file mode 100644 index 28381c1..0000000 --- a/src/SoftwarePipelining/SPMVE.ml +++ /dev/null @@ -1,299 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open SPBasic -open SPIMS - -module Mult = Map.Make (struct type t = reg let compare = compare end) - -let size_of_map1 m = - NI.fold (fun key v size -> size + 1) m 0 - -let size_of_map2 m = - Mult.fold (fun key v size -> size + 1) m 0 - -let sched_max_time sched = - NI.fold (fun node time max -> - if time > max then time else max - ) sched 0 - -let print_table t s = - Printf.fprintf SPDebug.dc "%s : \n" s; - let string_of_node_ev node = - match node with - | Some node -> string_of_node node - | None -> "_" - in - Array.iteri (fun i node -> Printf.fprintf SPDebug.dc "%i :: %s \n" i (string_of_node_ev node)) t - -let durations ddg sched ii = - - G.fold_vertex (fun node mult -> - match written node with - | None -> mult - | Some r -> - let raw_succs = get_succs_raw ddg node in - let durations = List.map (fun succ -> - let d = NI.find succ sched - NI.find node sched in - if d >= 0 then d - else ((sched_max_time sched - NI.find node sched) + NI.find succ sched) - ) raw_succs in - let duration = List.fold_left (fun max e -> if max > e then max else e) 0 durations in - Mult.add r ((duration / ii) + 1) mult (* cette division est surement fdausse*) - ) ddg Mult.empty - -let minimizer qi ur = - - let rec fill n = - if n <= ur then n :: fill (n + 1) else [] - in - - let l = fill qi in - let l = List.map (fun e -> (e,ur mod e)) l in - let l = List.filter (fun e -> snd e = 0) l in - let l = List.map fst l in - List.fold_left (fun min e -> if min < e then min else e) ur l - -let multiplicity ddg sched ii = - let qs = durations ddg sched ii in - (* Printf.printf "Quantite de variables necessaires : \n"; *) - (* Mult.iter (fun key mu -> print_reg key ; Printf.printf " |-> %i\n" mu) qs; *) - let unroll = Mult.fold (fun reg mult max -> - if mult > max then mult else max - ) qs 0 in - let mult = Mult.fold (fun reg mult mult -> - Mult.add reg (minimizer (Mult.find reg qs) unroll) mult - ) qs Mult.empty - in - (mult,unroll) - -let mve_kernel t ddg sched ii mult unroll = - - let regs = Array.make ii (fresh_regs 0) in - for i = 0 to (ii - 1) do - let fregs = fresh_regs unroll in - Array.iter print_reg fregs; - Printf.fprintf SPDebug.dc "\n"; - Array.set regs i fregs - done; - - let used_regs = ref [] in - - let index r i = Mult.find r mult - ( ((i / ii) + 1) mod Mult.find r mult) in - (* let pos i node inst = *) - (* let separation = *) - (* let b= NI.find node sched - NI.find inst sched in *) - (* if b >= 0 then b *) - (* else ((sched_max_time sched - NI.find inst sched) + NI.find node sched) *) - (* in *) - (* (i + separation) mod (ii * unroll) in *) - - let new_t = Array.copy t in - - for i = 0 to (Array.length t - 1) do - (* print_table new_t "Nouvelle table"; *) - match t.(i),new_t.(i) with - | Some insti, Some insti_mod -> - begin - match written insti with - | None -> () - | Some r -> - begin - let new_reg = - if Mult.find r mult = 0 then r - else regs.(i mod ii).(index r i - 1) in - if (not (List.mem (r,new_reg) !used_regs)) then used_regs := (r,new_reg) :: !used_regs; - let inst = reforge_writes insti_mod new_reg in - Printf.fprintf SPDebug.dc "Reecriture : %s --> %s \n" (string_of_node insti) (string_of_node inst); - Array.set new_t i (Some inst); - let succs = get_succs_raw ddg insti in - List.iter (fun node -> - - (* let lifetime = *) - (* let d = NI.find node sched - NI.find insti sched in *) - (* if d >= 0 then d *) - (* else ((sched_max_time sched - NI.find insti sched) + 1 + NI.find node sched) *) - (* in *) - - - (* ___Version 1___ *) - (* let lifetime = lifetime / ii in *) - (* let schedtime = *) - (* ((NI.find node sched) mod ii) (\* Position dans le premier bloc *\) *) - (* + (ii * (i / ii)) (\* Commencement du bloc ou on travail *\) *) - (* + (ii * lifetime ) (\* Decalage (Mult.find r mult - 1) *\) *) - (* + ((if (NI.find node sched mod ii) < (NI.find insti sched mod ii) then 0 else 0) * ii) *) - (* in *) - - (* Printf.printf "seed = %i - bloc = %i - slide = %i - corr = %i - id = %i \n" *) - (* ((NI.find node sched) mod ii) *) - (* (ii * (i / ii)) (ii * lifetime) *) - (* ((if (NI.find node sched mod ii) < (NI.find insti sched mod ii) then 1 else 0 ) * ii) *) - (* id; *) - (* Printf.printf "Successeur a traiter : %s ooo %i ooo\n" (string_of_node node) (Mult.find r mult); *) - - (* ___Version 2___ *) - let schedtime = - if (NI.find node sched > NI.find insti sched) - then i + (NI.find node sched - NI.find insti sched) - else i - NI.find insti sched + ii + NI.find node sched - (* let delta = NI.find insti sched - NI.find node sched in *) - (* (i - delta) + (((delta / ii) + 1) * ii) (\* (i - delta) + ii*\) *) - in - - (* then *) - - Printf.fprintf SPDebug.dc "i = %i - def = %i - use = %i - time = %i \n" - i (NI.find insti sched) (NI.find node sched) schedtime; - - - (* let id = pos i node insti in *) - let id = schedtime mod (unroll * ii) (* le tout modulo la tabl *) in - let id = (id + (unroll * ii)) mod (unroll * ii) in - Printf.fprintf SPDebug.dc "Positions to treat : %i \n" id; - let insttt = match new_t.(id) with - | Some inst -> inst - | None -> failwith "There should be an instruction" - in - let inst = reforge_reads insttt r new_reg in - Array.set new_t id (Some inst) - ) succs - end - end - | None, _ -> () - | _, None -> failwith "MVE : qqch mal compris" - done; - new_t,!used_regs - -let crunch_and_unroll sched ii ur = - let steady_s = Array.make ii None in - NI.iter (fun inst time -> - Array.set steady_s (time mod ii) (Some inst) - ) sched; - (* Printf.printf "Etat stable : \n"; *) - (* let string_of_node_ev node = *) - (* match node with *) - (* | Some node -> string_of_node node *) - (* | None -> "_" *) - (* in *) - (* Array.iteri (fun i node -> Printf.printf "%i :: %s \n" i (string_of_node_ev node)) steady_s; *) - let steady = Array.make (ii * ur) None in - for i = 0 to (ur - 1) do - for time = 0 to (ii - 1) do - Array.set steady (time + i * ii) (steady_s.(time)) - done - done; - steady - -let compute_iteration_table sched ii = - let t = Array.make ii None in - NI.iter (fun node time -> - Array.set t (NI.find node sched mod ii) (Some ((NI.find node sched / ii) + 1)) - ) sched; - t - -let compute_prolog steady min ii unroll schedule it = - - let prolog = ref [] in - let prolog_piece = ref [] in - - for i = (min - 1) downto 0 do - - let index = ((ii * (unroll - (min - i)))) mod (unroll * ii) in - prolog_piece := []; - - for j = 0 to (ii - 1) do (* copie du sous tableau *) - (* Printf.printf "i : %i - j : %i - index : %i \n" i j index; *) - match steady.(index + j), it.(j) with - | Some inst , Some iter -> - if iter <= (i + 1) then prolog_piece := inst :: !prolog_piece; (* i + 1 au lieu de i *) - | None, _ -> () - | _, _ -> failwith "compute_prolog: quelquechose est mal compris" - done; - - prolog := List.rev (!prolog_piece) @ !prolog - done; - - !prolog - - -let compute_epilog steady min ii unroll schedule it = - - let epilog = ref [] in - - for i = 0 to (min - 1) do - let index = (i mod unroll) * ii in - for j = 0 to (ii - 1) do - match steady.(index + j), it.(j) with - | Some inst , Some iter -> - if iter > (i + 1) then epilog := inst :: !epilog; - | None, _ -> () - | _, _ -> failwith "compute_prolog: quelquechose est mal compris" - done; - done; - List.rev (!epilog) - -let entrance = List.map (fun (a,b) -> (b,a)) - -let way_out prolog epilog used_regs = - let l = List.rev (prolog @ epilog) in - - let rec way_out_rec l wo = - match l with - | [] -> wo - | i :: l -> - begin - match written i with - | Some r -> - let mov = List.find (fun (a,b) -> r = b) used_regs in - if List.mem mov wo - then way_out_rec l wo - else way_out_rec l (mov :: wo) - | None -> way_out_rec l wo - end - in - - way_out_rec l [] - -let mve ddg sched ii = - assert (size_of_map1 sched = G.nb_vertex ddg); - Printf.fprintf SPDebug.dc "L'intervalle d'initiation est de : %i \n" ii; - Printf.fprintf SPDebug.dc "L'ordonnancement est le suivant : \n"; - print_schedule sched; - let (mult,unroll) = multiplicity ddg sched ii in - let unroll = unroll in (* changer pour tester, default doit etre egal a unroll *) - Printf.fprintf SPDebug.dc "Table de multiplicite : \n"; - Mult.iter (fun key mu -> print_reg key ; Printf.fprintf SPDebug.dc " |-> %i\n" mu) mult; - Printf.fprintf SPDebug.dc "Taux de deroulement de : %i \n" unroll; - let steady_state = crunch_and_unroll sched ii unroll in - let (steady_state,used_regs) = mve_kernel steady_state ddg sched ii mult unroll in - print_table steady_state "Table finale"; - let min = ((sched_max_time sched) / ii) + 1 in - Printf.fprintf SPDebug.dc "min : %i \n" min; - let iteration_table = compute_iteration_table sched ii in - Printf.fprintf SPDebug.dc "Table d'iteration \n"; - Array.iteri (fun i elt -> - match elt with - | Some elt -> - Printf.fprintf SPDebug.dc "%i : %i\n" i elt - | None -> Printf.fprintf SPDebug.dc "%i : _ \n" i - ) iteration_table; - let prolog = compute_prolog steady_state min ii unroll sched iteration_table in - let prolog = List.filter (fun e -> not (is_cond e)) prolog in - let epilog = compute_epilog steady_state min ii unroll sched iteration_table in - let epilog = List.filter (fun e -> not (is_cond e)) epilog in - Printf.fprintf SPDebug.dc "Prologue: \n"; - List.iter (fun node -> Printf.fprintf SPDebug.dc "%s \n" (string_of_node node)) prolog; - Printf.fprintf SPDebug.dc "Epilogue: \n"; - List.iter (fun node -> Printf.fprintf SPDebug.dc "%s \n" (string_of_node node)) epilog; - let way_out = way_out prolog epilog used_regs in - (steady_state,prolog,epilog,min - 1,unroll,entrance used_regs, way_out) diff --git a/src/SoftwarePipelining/SPMVE.mli b/src/SoftwarePipelining/SPMVE.mli deleted file mode 100644 index 2da367d..0000000 --- a/src/SoftwarePipelining/SPMVE.mli +++ /dev/null @@ -1,17 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open SPBasic -open SPIMS - -val mve : G.t -> int NI.t -> int -> - (G.V.t option) array * G.V.t list * G.V.t list * int * int * (reg * reg) list * (reg * reg) list diff --git a/src/SoftwarePipelining/SPSymbolic_evaluation.ml b/src/SoftwarePipelining/SPSymbolic_evaluation.ml deleted file mode 100644 index c99afc8..0000000 --- a/src/SoftwarePipelining/SPSymbolic_evaluation.ml +++ /dev/null @@ -1,226 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open Registers -open Op -open AST -open SPBase_types -open Camlcoq - -type symbolic_value = - | Sreg of reg - | Sop of operation * symbolic_value list - | Sload of memory_chunk * addressing * symbolic_value list * symbolic_mem - -and symbolic_mem = - | Smem - | Sstore of memory_chunk * addressing * symbolic_value list * symbolic_value * symbolic_mem - -module State = Map.Make (struct type t = reg let compare = compare end) - -module Cons = Set.Make (struct type t = symbolic_value let compare = compare end) - -type symbolic_state = symbolic_value State.t * Cons.t - -let initial_state = State.empty -let initial_mem = Smem -let initial_cons = Cons.empty - -exception Not_straight - -let find res st = - try State.find res st - with - | Not_found -> Sreg res - -let rec get_args st = function - | [] -> [] - | arg::args -> find arg st :: get_args st args - -let rec symbolic_evaluation st sm cs = function - | [] -> (st,sm,cs) - | Inop :: l -> symbolic_evaluation st sm cs l - - | Iop (Omove, [src], dst) :: l -> - symbolic_evaluation (State.add dst (find src st) st) sm cs l - - | Iop (op, args, dst) :: l -> - let sym_val = Sop (op,get_args st args) in - symbolic_evaluation (State.add dst sym_val st) sm (Cons.add sym_val cs) l - - | Iload (chunk, mode, args, dst) :: l -> - let sym_val = Sload (chunk, mode, get_args st args, sm) in - symbolic_evaluation (State.add dst sym_val st) sm (Cons.add sym_val cs) l - - | Istore (chunk, mode, args, src) :: l -> - let sym_mem = Sstore (chunk, mode, get_args st args, find src st, sm) in - symbolic_evaluation st sym_mem cs l - - | _ :: l -> raise Not_straight - -type osv = - | Oresource of resource - | Oop of operation - | Oload of memory_chunk * addressing - | Ostore of memory_chunk * addressing - -let string_of_osv = function - | Oresource (Reg r) -> Printf.sprintf "reg %i" (P.to_int r) - | Oresource Mem -> "mem" - | Oop op -> string_of_op op - | Oload (mc,addr) -> "load" - | Ostore (mc,addr) -> "store" - -type ident = int - -module S = Graph.Persistent.Digraph.Abstract - (struct type t = osv * ident end) - -let name_of_vertex v = - let (osv,id) = S.V.label v in - Printf.sprintf "%i" id - -let string_of_vertex v = - let (osv,_) = S.V.label v in - Printf.sprintf "%s" (string_of_osv osv) - -module DisplayTree = struct - include S - let vertex_name v = name_of_vertex v - let graph_attributes _ = [] - let default_vertex_attributes _ = [] - let vertex_attributes v = [`Label (string_of_vertex v)] - let default_edge_attributes _ = [] - let edge_attributes _ = [] - let get_subgraph _ = None -end -module DotTree = Graph.Graphviz.Dot(DisplayTree) - -let dot_output_ss g f = - let oc = open_out f in - DotTree.output_graph oc g; - close_out oc - -module Build = Graph.Builder.P (S) -module Op = Graph.Oper.Make (Build) - -let counter = ref 0 - -let rec convert_sv_rec sv graph = - incr counter; - match sv with - | Sreg res -> - let node = S.V.create (Oresource (Reg res), !counter) in - let graph = S.add_vertex graph node in - (graph,node) - - | Sop (op, svl) -> - let node = S.V.create (Oop op, !counter) in - let (graph, node_l) = List.fold_right (fun sv (graph,node_l) -> - let (graph,node) = convert_sv_rec sv graph in - graph, node :: node_l - ) svl (graph,[]) in - let graph = S.add_vertex graph node in - let graph = List.fold_right (fun n graph -> - S.add_edge graph node n - ) node_l graph in - (graph,node) - - - | Sload (mc,addr,svl,sm) -> - let node = S.V.create (Oload (mc, addr), !counter) in - let (graph, node_l) = List.fold_right (fun sv (graph,node_l) -> - let (graph,node) = convert_sv_rec sv graph in - graph, node :: node_l - ) svl (graph,[]) in - let (graph,node_m) = convert_sm_rec sm graph in - let graph = S.add_vertex graph node in - let graph = List.fold_right (fun n graph -> - S.add_edge graph node n - ) node_l graph in - let graph = S.add_edge graph node node_m in - (graph,node) - -and convert_sm_rec sm graph = - incr counter; - match sm with - | Smem -> - let node = S.V.create (Oresource Mem, !counter) in - let graph = S.add_vertex graph node in - (graph,node) - - | Sstore (mc,addr,svl,sv,sm) -> - let node = S.V.create (Ostore (mc, addr), !counter) in - let (graph, node_l) = List.fold_right (fun sv (graph,node_l) -> - let (graph,node) = convert_sv_rec sv graph in - graph, node :: node_l - ) svl (graph,[]) in - let (graph, n) = convert_sv_rec sv graph in - let (graph, node_m) = convert_sm_rec sm graph in - let graph = S.add_vertex graph node in - let graph = List.fold_right (fun n graph -> - S.add_edge graph node n - ) node_l graph in - let graph = S.add_edge graph node n in - let graph = S.add_edge graph node node_m in - (graph,node) - -let convert_sv sv = convert_sv_rec sv S.empty -let convert_sm sm = convert_sm_rec sm S.empty - -let convert_sym st sm regs = - let graph = State.fold (fun res sv g -> - if (not (List.mem res regs)) then g - else - let (graph,head) = convert_sv sv in - incr counter; - let src = S.V.create (Oresource (Reg res), !counter) in - let graph = S.add_vertex graph src in - let graph = S.add_edge graph src head in - Op.union g graph - ) st S.empty - in - let graph' = - let (graph,head) = convert_sm sm in - incr counter; - let src = S.V.create (Oresource Mem, !counter) in - let graph = S.add_vertex graph src in - let graph = S.add_edge graph src head in - graph - in - Op.union graph graph' - -let display_st name l regs = - let (st,sm,_) = symbolic_evaluation initial_state initial_mem initial_cons l in - let g = convert_sym st sm regs in - let addr = SPDebug.name ^ name in - dot_output_ss g addr ; - ignore (Sys.command ("(dot -Tpng " ^ addr ^ " -o " ^ addr ^ ".png ; rm -f " ^ addr ^ ") & ")) (* & *) - - -let symbolic_equivalence (st1,sm1,cs1) (st2,sm2,cs2) regs = - Printf.printf "|cs1| = %i - |cs2| = %i \n" (Cons.cardinal cs1) (Cons.cardinal cs2); - (List.fold_right (fun res b -> - find res st1 = find res st2 && b - ) regs true - && sm1 = sm2 - && Cons.equal cs1 cs2) - -let symbolic_evaluation = symbolic_evaluation initial_state initial_mem initial_cons - -let symbolic_condition i l = - match i with - | Icond (cond,args) -> - let args = to_caml_list args in - let (st,sm,cs) = symbolic_evaluation l in - (cond,List.map (fun r -> find r st) args) - | _ -> failwith "Not a condition !\n" diff --git a/src/SoftwarePipelining/SPTyping.ml b/src/SoftwarePipelining/SPTyping.ml deleted file mode 100644 index 9b9c679..0000000 --- a/src/SoftwarePipelining/SPTyping.ml +++ /dev/null @@ -1,526 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -(*open Datatypes -open List -open Camlcoq -open Maps -open AST -open Op -open Registers -open RTL - -open Conventions -open Coqlib -open Errors -open Specif - -exception Type_error of string - -let env = ref (PTree.empty : typ PTree.t) - -let set_type r ty = - match PTree.get r !env with - | None -> env := PTree.set r ty !env - | Some ty' -> if ty <> ty' then - begin - Printf.fprintf SPDebug.dc "Failed to type register : %i " (P.to_int r); - raise (Type_error "type mismatch") - end - -let rec set_types rl tyl = - match rl, tyl with - | [], [] -> () - | r1 :: rs, ty1 :: tys -> set_type r1 ty1; set_types rs tys - | _, _ -> raise (Type_error "arity mismatch") - -(* First pass: process constraints of the form typeof(r) = ty *) - -let type_instr retty (pc, i) = - Printf.fprintf SPDebug.dc "typage de l'instruction : %i \n" (P.to_int pc); - match i with - | Inop(_) -> - () - | Iop(Omove, _, _, _) -> - () - | Iop(op, args, res, _) -> - let (targs, tres) = type_of_operation op in - set_types args targs; set_type res tres - | Iload(chunk, addr, args, dst, _) -> - set_types args (type_of_addressing addr); - set_type dst (type_of_chunk chunk) - | Istore(chunk, addr, args, src, _) -> - set_types args (type_of_addressing addr); - set_type src (type_of_chunk chunk) - | Icall(sg, ros, args, res, _) -> - begin try - begin match ros with - | Coq_inl r -> set_type r Tint - | Coq_inr _ -> () - end; - set_types args sg.sig_args; - set_type res (match sg.sig_res with Tret t -> t | _ -> Tint); - with Type_error msg -> - let name = - match ros with - | Coq_inl _ -> "" - | Coq_inr id -> extern_atom id in - raise(Type_error (Printf.sprintf "type mismatch in Icall(%s): %s" - name msg)) - end - | Itailcall(sg, ros, args) -> - begin try - begin match ros with - | Coq_inl r -> set_type r Tint - | Coq_inr _ -> () - end; - set_types args sg.sig_args; - if sg.sig_res <> retty then - raise (Type_error "mismatch on return type") - with Type_error msg -> - let name = - match ros with - | Coq_inl _ -> "" - | Coq_inr id -> extern_atom id in - raise(Type_error (Printf.sprintf "type mismatch in Itailcall(%s): %s" - name msg)) - end -(* | Ialloc(arg, res, _) -> - set_type arg Tint; set_type res Tint*) - | Icond(cond, args, _, _) -> - set_types args (type_of_condition cond) - | Ireturn(optres) -> - begin match optres, retty with - | None, Tvoid -> () - | Some r, Tret ty -> set_type r ty - | _, _ -> raise (Type_error "type mismatch in Ireturn") - end - -let type_pass1 retty instrs = - List.iter (type_instr retty) instrs - -(* Second pass: extract move constraints typeof(r1) = typeof(r2) - and solve them iteratively *) - -let rec extract_moves = function - | [] -> [] - | (pc, i) :: rem -> - match i with - | Iop(Omove, [r1], r2, _) -> - (r1, r2) :: extract_moves rem - | Iop(Omove, _, _, _) -> - raise (Type_error "wrong Omove") - | _ -> - extract_moves rem - -let changed = ref false - -let rec solve_moves = function - | [] -> [] - | (r1, r2) :: rem -> - match (PTree.get r1 !env, PTree.get r2 !env) with - | Some ty1, Some ty2 -> - if ty1 = ty2 - then (changed := true; solve_moves rem) - else raise (Type_error "type mismatch in Omove") - | Some ty1, None -> - env := PTree.set r2 ty1 !env; changed := true; solve_moves rem - | None, Some ty2 -> - env := PTree.set r1 ty2 !env; changed := true; solve_moves rem - | None, None -> - (r1, r2) :: solve_moves rem - -let rec iter_solve_moves mvs = - changed := false; - let mvs' = solve_moves mvs in - if !changed then iter_solve_moves mvs' - -let type_pass2 instrs = - iter_solve_moves (extract_moves instrs) - -let typeof e r = - match PTree.get r e with Some ty -> ty | None -> Tint - -let infer_type_environment f instrs = - try - env := PTree.empty; - set_types f.fn_params f.fn_sig.sig_args; - type_pass1 f.fn_sig.sig_res instrs; - type_pass2 instrs; - let e = !env in - env := PTree.empty; - Some(typeof e) - with Type_error msg -> - Printf.eprintf "Error during RTL type inference: %s\n" msg; - None - -(** val typ_eq : typ -> typ -> bool **) - -let typ_eq t1 t2 = - match t1 with - | Tint -> (match t2 with - | Tint -> true - | Tfloat -> false) - | Tfloat -> (match t2 with - | Tint -> false - | Tfloat -> true) - -(** val opt_typ_eq : typ option -> typ option -> bool **) - -let opt_typ_eq t1 t2 = - match t1 with - | Some x -> (match t2 with - | Some t -> typ_eq x t - | None -> false) - | None -> (match t2 with - | Some t -> false - | None -> true) - -(** val check_reg : regenv -> reg -> typ -> bool **) - -let check_reg env r ty = - match typ_eq (env r) ty with - | true -> true - | false -> false - -(** val check_regs : regenv -> reg list -> typ list -> bool **) - -let rec check_regs env rl tyl = - match rl with - | [] -> - (match tyl with - | [] -> true - | (t :: l) -> false) - | r1 :: rs -> - (match tyl with - | [] -> false - | (ty :: tys) -> - (match typ_eq (env r1) ty with - | true -> check_regs env rs tys - | false -> false)) - -(** val check_op : regenv -> operation -> reg list -> reg -> bool **) - -let check_op env op args res0 = - let (targs, tres) = type_of_operation op in - (match check_regs env args targs with - | true -> - (match typ_eq (env res0) tres with - | true -> true - | false -> false) - | false -> false) - -(** val check_successor : coq_function -> node -> bool **) - -let check_successor funct s = - match Maps.PTree.get s funct.fn_code with - | Some i -> true - | None -> false - -(** val check_instr : coq_function -> regenv -> instruction -> bool **) - -let check_instr funct env = function - | Inop s -> check_successor funct s - | Iop (op, args, res0, s) -> - (match op with - | Omove -> - (match args with - | Coq_nil -> false - | Coq_cons (arg, l) -> - (match l with - | Coq_nil -> - (match typ_eq (env arg) (env res0) with - | true -> check_successor funct s - | false -> false) - | Coq_cons (r, l0) -> false)) - | Ointconst i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ofloatconst f -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) -(* | Oaddrsymbol (i0, i1) -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false)*) -(* | Oaddrstack i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false)*) - | Ocast8signed -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ocast8unsigned -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ocast16signed -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ocast16unsigned -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) -(* | Oadd -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oaddimm i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false)*) - | Osub -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Osubimm i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Omul -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Omulimm i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Odiv -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Odivu -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oand -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oandimm i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oor -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oorimm i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oxor -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oxorimm i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Onand -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Onor -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Onxor -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oshl -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oshr -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oshrimm i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oshrximm i0 -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oshru -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Orolm (i0, i1) -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Onegf -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oabsf -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Oaddf -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Osubf -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Omulf -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Odivf -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Omuladdf -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Omulsubf -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Osingleoffloat -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ointoffloat -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ointuoffloat -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ofloatofint -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ofloatofintu -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false) - | Ocmp c -> - (match check_op env op args res0 with - | true -> check_successor funct s - | false -> false)) - | Iload (chunk, addr, args, dst, s) -> - (match check_regs env args (type_of_addressing addr) with - | true -> - (match typ_eq (env dst) (type_of_chunk chunk) with - | true -> check_successor funct s - | false -> false) - | false -> false) - | Istore (chunk, addr, args, src, s) -> - (match check_regs env args (type_of_addressing addr) with - | true -> - (match typ_eq (env src) (type_of_chunk chunk) with - | true -> check_successor funct s - | false -> false) - | false -> false) - | Icall (sig0, ros, args, res0, s) -> - (match match ros with - | Coq_inl r -> - (match typ_eq (env r) Tint with - | true -> check_regs env args sig0.sig_args - | false -> false) - | Coq_inr s0 -> check_regs env args sig0.sig_args with - | true -> - (match typ_eq (env res0) (proj_sig_res sig0) with - | true -> check_successor funct s - | false -> false) - | false -> false) - | Itailcall (sig0, ros, args) -> - (match match match ros with - | Coq_inl r -> - (match typ_eq (env r) Tint with - | true -> check_regs env args sig0.sig_args - | false -> false) - | Coq_inr s -> check_regs env args sig0.sig_args with - | true -> - proj_sumbool - (opt_typ_eq sig0.sig_res funct.fn_sig.sig_res) - | false -> false with - | true -> tailcall_is_possible sig0 - | false -> false) -(* | Ialloc (arg, res0, s) -> - (match typ_eq (env arg) Tint with - | true -> - (match typ_eq (env res0) Tint with - | true -> check_successor funct s - | false -> false) - | false -> false)*) - | Icond (cond, args, s1, s2) -> - (match match check_regs env args (type_of_condition cond) with - | true -> check_successor funct s1 - | false -> false with - | true -> check_successor funct s2 - | false -> false) - | Ireturn optres -> - (match optres with - | Some r -> - (match funct.fn_sig.sig_res with - | Some t -> - (match typ_eq (env r) t with - | true -> true - | false -> false) - | None -> false) - | None -> - (match funct.fn_sig.sig_res with - | Some t -> false - | None -> true)) - -(** val check_params_norepet : reg list -> bool **) - -let check_params_norepet params = - match list_norepet_dec Registers.Reg.eq params with - | true -> true - | false -> false - -(** val check_instrs : coq_function -> regenv -> (node, instruction) prod - list -> bool **) - -let rec check_instrs funct env = function - | Coq_nil -> true - | Coq_cons (p, rem) -> - let Coq_pair (pc, i) = p in - (match check_instr funct env i with - | true -> check_instrs funct env rem - | false -> false) - -(** val type_function : coq_function -> unit **) - -let type_function f = - let instrs = Maps.PTree.elements f.fn_code in - match infer_type_environment f instrs with - | Some env -> - (match match match match check_regs env f.fn_params - f.fn_sig.sig_args with - | true -> check_params_norepet f.fn_params - | false -> false with - | true -> check_instrs f env instrs - | false -> false with - | true -> check_successor f f.fn_entrypoint - | false -> false with - | true -> Printf.fprintf SPDebug.dc "The code is well typed\n" - | false -> failwith "Type checking failure\n") - | None -> failwith "Type inference failure\n" -*) diff --git a/src/SoftwarePipelining/SPTyping.mli b/src/SoftwarePipelining/SPTyping.mli deleted file mode 100644 index dd27875..0000000 --- a/src/SoftwarePipelining/SPTyping.mli +++ /dev/null @@ -1,14 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - - -(*val type_function : RTL.coq_function -> unit*) diff --git a/src/SoftwarePipelining/SoftwarePipelining.ml b/src/SoftwarePipelining/SoftwarePipelining.ml deleted file mode 100644 index 0ba6d9d..0000000 --- a/src/SoftwarePipelining/SoftwarePipelining.ml +++ /dev/null @@ -1,74 +0,0 @@ -(***********************************************************************) -(* *) -(* Compcert Extensions *) -(* *) -(* Jean-Baptiste Tristan *) -(* *) -(* All rights reserved. This file is distributed under the terms *) -(* described in file ../../LICENSE. *) -(* *) -(***********************************************************************) - - -open SPBasic -open SPIMS -open SPMVE -open RTL - -let clean t = - - let rec clean_rec i = - match i with - | 0 -> [] - | n -> - begin - match t.(i - 1) with - | None -> clean_rec (i - 1) - | Some inst -> inst :: clean_rec (i - 1) - end - in - let l = List.rev (clean_rec (Array.length t)) in - List.hd l :: (List.filter (fun e -> not (is_cond e)) (List.tl l)) - -let print_nodes = List.iter (fun n -> Printf.printf "%s \n" (string_of_node n)) - -(* random heuristic *) - -let find node schedule opt = - try NI.find node schedule with - | Not_found -> opt - -(* A random heuristic is used to pick the next instruction to be scheduled from the unscheduled - * instructions. The scheduled instructions are given to the function, and the unscheduled - * instructions are created by taking all the instructions that are not in the scheduled list. - *) -let random ddg schedule = - let unscheduled = G.fold_vertex (fun node l -> - match find node schedule None with - | Some v -> l - | None -> node :: l - ) ddg [] in - let bound = List.length unscheduled in - Random.self_init (); - List.nth unscheduled (Random.int bound) - -(* tought heuristics *) - -module Topo = Graph.Topological.Make (G) -module Scc = Graph.Components.Make (G) - -let order = ref [] - -let pipeliner ddg = - order := List.flatten (Scc.scc_list ddg); - let (sched,ii) = SPIMS.pipeliner ddg random in - let (steady,prolog,epilog,min,unroll,entrance,way_out) = SPMVE.mve ddg sched ii in - let steady_state = clean steady in - if min <= 0 then None - else - Some {steady_state = steady_state; prolog = prolog; epilog = epilog; min = min; unrolling = unroll; - ramp_up = entrance; ramp_down = way_out} - - -let pipeline f = - SPBasic.apply_pipeliner f pipeliner ~debug:false diff --git a/src/pipelining/LICENSE b/src/pipelining/LICENSE new file mode 100644 index 0000000..e275fa0 --- /dev/null +++ b/src/pipelining/LICENSE @@ -0,0 +1,19 @@ +Copyright (c) 2008,2009,2010 Jean-Baptiste Tristan and INRIA + +Permission is hereby granted, free of charge, to any person obtaining a copy +of this software and associated documentation files (the "Software"), to deal +in the Software without restriction, including without limitation the rights +to use, copy, modify, merge, publish, distribute, sublicense, and/or sell +copies of the Software, and to permit persons to whom the Software is +furnished to do so, subject to the following conditions: + +The above copyright notice and this permission notice shall be included in +all copies or substantial portions of the Software. + +THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR +IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, +FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE +AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER +LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, +OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE +SOFTWARE. diff --git a/src/pipelining/SPBase_types.ml b/src/pipelining/SPBase_types.ml new file mode 100644 index 0000000..ba92340 --- /dev/null +++ b/src/pipelining/SPBase_types.ml @@ -0,0 +1,128 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open Camlcoq +open Op +open Registers +open Memory +open Mem +open AST + +type ('a,'b) sum = ('a,'b) Datatypes.sum + +type instruction = + | Inop + | Iop of operation * reg list * reg + | Iload of memory_chunk * addressing * reg list * reg + | Istore of memory_chunk * addressing * reg list * reg + | Icall of signature * (reg, ident) sum * reg list * reg + | Itailcall of signature * (reg, ident) sum * reg list + | Icond of condition * reg list + | Ireturn of reg option + +type resource = Reg of reg | Mem + +let inst_coq_to_caml = function + | RTL.Inop succ -> Inop + | RTL.Iop (op, args, dst, succ) -> Iop (op, args, dst) + | RTL.Iload (chunk, mode, args, dst, succ) -> Iload (chunk, mode, args, dst) + | RTL.Istore (chunk, mode, args, src, succ) -> Istore (chunk, mode, args, src) + | RTL.Icall (sign, id, args, dst, succ) -> Icall (sign, id, args, dst) + | RTL.Itailcall (sign, id, args) -> Itailcall (sign, id, args) + | RTL.Icond (cond, args, succ1, succ2) -> Icond (cond, args) + | RTL.Ireturn (res) -> Ireturn (res) + +let inst_caml_to_coq i (links : RTL.node list) = + match i,links with + | Inop,[p] -> RTL.Inop p + | Iop (op, args, dst),[p] -> RTL.Iop (op, args, dst,p) + | Iload (chunk, mode, args, dst),[p] -> RTL.Iload (chunk, mode, args,dst,p) + | Istore (chunk, mode, args, src),[p] -> RTL.Istore (chunk, mode, args, src,p) + | Icall (sign, id, args, dst),[p] -> RTL.Icall (sign, id, args, dst,p) + | Itailcall (sign, id, args),[] -> RTL.Itailcall (sign, id, args) + | Icond (cond, args),[p1;p2] -> RTL.Icond (cond,args,p1,p2) + | Ireturn (res),[] -> RTL.Ireturn res + | _,_ -> failwith "Echec lors de la conversion des instrucitons internes vers compcert" + + +let print_inst node = string_of_int (snd node) + +let to_int = fun n -> P.to_int n +let to_binpos = fun n -> P.of_int n + +let rec string_of_args args = + match args with + | [] -> "" + | arg :: args -> "r" ^ (string_of_int (to_int arg)) ^ " " ^ string_of_args args + +let string_of_z z = string_of_int (Z.to_int z) + +let string_of_comparison = function + | Integers.Ceq -> "eq" + | Integers.Cne -> "ne" + | Integers.Clt -> "lt" + | Integers.Cle -> "le" + | Integers.Cgt -> "gt" + | Integers.Cge -> "ge" + +let string_of_cond = function + | Ccomp c -> Printf.sprintf "comp %s" (string_of_comparison c) + | Ccompu c -> Printf.sprintf "compu %s" (string_of_comparison c) + | Ccompimm (c,i) -> Printf.sprintf "compimm %s %s" (string_of_comparison c) (string_of_z i) + | Ccompuimm (c,i) -> Printf.sprintf "compuimm %s %s" (string_of_comparison c) (string_of_z i) + | Ccompf c -> Printf.sprintf "compf %s" (string_of_comparison c) + | Cnotcompf c -> Printf.sprintf "notcompf %s" (string_of_comparison c) + | Cmaskzero i -> Printf.sprintf "maskzero %s" (string_of_z i) + | Cmasknotzero i -> Printf.sprintf "masknotzero %s" (string_of_z i) + +let string_of_op = function + | Omove -> "move" + | Ointconst i -> Printf.sprintf "intconst %s" (string_of_z i) + | Ofloatconst f -> Printf.sprintf "floatconst %s" (string_of_float (camlfloat_of_coqfloat32 f)) + | Ocast8signed -> "cast8signed" + | Ocast8unsigned -> "cast8unsigned" + | Ocast16signed -> "cast16signed" + | Ocast16unsigned -> "cast16unsigned" + | Osub -> "sub" + | Omul -> "mul" + | Omulimm i -> Printf.sprintf "mulimm %s" (string_of_z i) + | Odiv -> "div" + | Odivu -> "divu" + | Oand -> "and" + | Oandimm i -> Printf.sprintf "andimm %s" (string_of_z i) + | Oor -> "or" + | Oorimm i -> Printf.sprintf "orimm %s" (string_of_z i) + | Oxor -> "xor" + | Oxorimm i -> Printf.sprintf "xorimm %s" (string_of_z i) + | Oshl -> "shl" + | Oshr -> "shr" + | Oshrimm i -> Printf.sprintf "shrimm %s" (string_of_z i) + | Oshrximm i -> Printf.sprintf "shrximm %s" (string_of_z i) + | Oshru -> "shru" + | Onegf -> "negf" + | Oabsf -> "absf" + | Oaddf -> "addf" + | Osubf -> "subf" + | Omulf -> "mulf" + | Odivf -> "divf" + | Osingleoffloat -> "singleoffloat" + | Ofloatofint -> "floatofint" + | Ocmp c -> Printf.sprintf "cmpcmpcmp" + | Olea _ -> "lea" + +let to_coq_list = function + | [] -> [] + | e :: l -> e :: l + +let to_caml_list = function + | [] -> [] + | e :: l -> e :: l diff --git a/src/pipelining/SPBasic.ml b/src/pipelining/SPBasic.ml new file mode 100644 index 0000000..32234b8 --- /dev/null +++ b/src/pipelining/SPBasic.ml @@ -0,0 +1,819 @@ +(***********************************************************************) + (* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open RTL +open Camlcoq +open Graph.Pack.Digraph +open Op +open Registers +open Memory +open Mem +open AST +open SPBase_types +open SPSymbolic_evaluation + +type node = instruction * int +module G = Graph.Persistent.Digraph.AbstractLabeled + (struct type t = node end) + (struct type t = int let compare = compare let default = 0 end) +module Topo = Graph.Topological.Make (G) +module Dom = Graph.Dominator.Make (G) +module Index = Map.Make (struct type t = int let compare = compare end) + +let string_of_instruction node = + match G.V.label node with + | (Inop,n) -> Printf.sprintf "%i nop" n + | (Iop (op, args, dst),n) -> Printf.sprintf "%i r%i := %s %s" n (to_int dst) (string_of_op op) (string_of_args args) + | (Iload (chunk, mode, args, dst),n) -> Printf.sprintf "%i r%i := load %s" n (to_int dst) (string_of_args args) + | (Istore (chunk, mode, args, src),n) -> Printf.sprintf "%i store %i %s" n (to_int src) (string_of_args args) + | (Icall (sign, id, args, dst),n) -> Printf.sprintf "%i call" n + | (Itailcall (sign, id, args),n) -> Printf.sprintf "%i tailcall" n + (* | (Ialloc (dst, size),n) -> Printf.sprintf "%i %i := alloc" n (to_int dst) *) + | (Icond (cond, args),n) -> Printf.sprintf "%i cond %s %s" n (string_of_cond cond) (string_of_args args) + | (Ireturn (res),n) -> Printf.sprintf "%i return" n + +let string_of_node = string_of_instruction + +module Display = struct + include G + let vertex_name v = print_inst (V.label v) + let graph_attributes _ = [] + let default_vertex_attributes _ = [] + let vertex_attributes v = [`Label (string_of_instruction v)] + let default_edge_attributes _ = [] + let edge_attributes e = [`Label (string_of_int (G.E.label e))] + let get_subgraph _ = None +end +module Dot_ = Graph.Graphviz.Dot(Display) + +let dot_output g f = + let oc = open_out f in + Dot_.output_graph oc g; + close_out oc + +let display g name = + let addr = SPDebug.name ^ name in + dot_output g addr ; + ignore (Sys.command ("(dot -Tpng " ^ addr ^ " -o " ^ addr ^ ".png ; rm -f " ^ addr ^ ") & ")) + +(******************************************) + +type cfg = {graph : G.t; start : G.V.t} + +(* convert traduit un graphe RTL compcert en un graphe RTL ocamlgraph*) + +let convert f = + + let make_node inst key = + let inst = inst_coq_to_caml inst in + G.V.create (inst, to_int key) + in + + let (graph, index) = Maps.PTree.fold (fun (g,m) key inst -> + let node = make_node inst key in + (G.add_vertex g node, Index.add (to_int key) node m) + ) f.fn_code (G.empty,Index.empty) + in + + let succ = RTL.successors_map f in + let rec link n succs g = + match succs with + | [] -> g + | pos::[] -> + G.add_edge g (Index.find (to_int n) index) (Index.find (to_int pos) index) + | pos1::pos2::[] -> + let g = G.add_edge_e g (G.E.create (Index.find (to_int n) index) 1 (Index.find (to_int pos1) index)) in + G.add_edge_e g (G.E.create (Index.find (to_int n) index) 2 (Index.find (to_int pos2) index)) + | _ -> failwith "convert : trop de successeurs" + + in + + let graph = Maps.PTree.fold ( fun g key inst -> + link key (match Maps.PTree.get key succ with + Some x -> x | _ -> failwith "Could not index") g + ) f.fn_code graph + in + + {graph = graph; start = Index.find (to_int (f.fn_entrypoint)) index} + + +let convert_back g = + + G.fold_vertex (fun node m -> + let v = G.V.label node in + match (fst v) with + | Icond (_,_) -> + begin + let l = + match G.succ_e g node with + | [e1;e2] -> + if G.E.label e1 > G.E.label e2 + then [G.E.dst e2;G.E.dst e1] + else [G.E.dst e1;G.E.dst e2] + | _ -> failwith "convert_back: nombre de successeurs incoherent" + in + let succs = List.map (fun s -> to_binpos (snd (G.V.label s))) l in + Maps.PTree.set (to_binpos (snd v)) (inst_caml_to_coq (fst v) succs) m + end + | _ -> + let succs = List.map (fun s -> to_binpos (snd (G.V.label s))) (G.succ g node) in + Maps.PTree.set (to_binpos (snd v)) (inst_caml_to_coq (fst v) succs) m + ) g Maps.PTree.empty + + + + +(* dominator_tree calcule l'arbre de domination grace au code de FP *) +let dominator_tree f = + Dom.compute_idom f.graph f.start + +(* detect_loops, find the loops in the graph and returns the list of nodes in it, + in dominating order !!! This is of great importance, we suppose that it is ordered + when we build the dependency graph *) +let detect_loops graph dom_tree = + let rec is_dominating v1 v2 l = (* does v1 dominate v2 *) + match dom_tree v2 with + | v -> if v1 = v then Some (v :: l) + else is_dominating v1 v (v :: l) + | exception (Not_found) -> None + in + + G.fold_edges (fun v1 v2 loops -> + match is_dominating v2 v1 [v1] with + | None -> loops + | Some loop -> (v2,loop) :: loops + ) graph [] + +let print_index node = + Printf.printf "%i " (snd (G.V.label node)) + +let print_instruction node = + match G.V.label node with + | (Inop,n) -> Printf.printf "%i : Inop \n" n + | (Iop (op, args, dst),n) -> Printf.printf "%i : Iop \n" n + | (Iload (chunk, mode, args, dst),n) -> Printf.printf "%i : Iload \n" n + | (Istore (chunk, mode, args, src),n) -> Printf.printf "%i : Istore \n" n + | (Icall (sign, id, args, dst),n) -> Printf.printf "%i : Icall \n" n + | (Itailcall (sign, id, args),n) -> Printf.printf "%i : Itailcall \n" n + (*| (Ialloc (dst, size),n) -> Printf.printf "%i : Ialloc \n" n *) + | (Icond (cond, args),n) -> Printf.printf "%i : Icond \n" n + | (Ireturn (res),n) -> Printf.printf "%i : Ireturn \n" n + +let is_rewritten node r = + match fst (G.V.label node) with + | Inop -> false + | Iop (op, args, dst) -> if dst = r then true else false + | Iload (chunk, mode, args, dst) -> if dst = r then failwith "J'ai degote une boucle ZARBI !!!" else false + | Istore (chunk, mode, args, src) -> false + | Icall (sign, id, args, dst) -> failwith "call in a loop" + | Itailcall (sign, id, args) -> failwith "tailcall in a loop" + (* | Ialloc (dst, size) -> if dst = r then failwith "J'ai degote une boucle ZARBI !!!" else false *) + | Icond (cond, args) -> false + | Ireturn (res) -> failwith "return in a loop" + +let is_variant r loop = + List.fold_right (fun node b -> + is_rewritten node r || b + ) loop false + + +let is_pipelinable loop = (* true if loop is innermost and without control *) + + let is_acceptable node = + match fst (G.V.label node) with + | Icall _ | Itailcall _ | Ireturn _ | Icond _ (*| Ialloc _*) | Iop ((Ocmp _),_,_)-> false + | _ -> true + in + + let is_branching node = + match fst (G.V.label node) with + | Icond _ -> true + | _ -> false + in + + let is_nop node = + match fst (G.V.label node) with + | Inop -> true + | _ -> false + in + + let is_OK_aux l = + List.fold_right (fun n b -> is_acceptable n && b) l true + in + + let is_bounded node loop = + match G.V.label node with + | (Icond (cond, args),n) -> + let args = to_caml_list args in + begin + match args with + | [] -> false + | r :: [] -> is_variant r loop (* used to be not *) + | r1 :: r2 :: [] -> + begin + match is_variant r1 loop, is_variant r2 loop with + | true, true -> false + | false, true -> true + | true, false -> true + | false, false -> false + end + | _ -> false + end + | _ -> false + in + + match List.rev loop with + | v2 :: v1 :: l -> ((*Printf.printf "is_nop: %s | " (is_nop v1 |> string_of_bool);*) + Printf.printf "is_branching: %s | " (is_branching v2 |> string_of_bool); + Printf.printf "is_OK_aux: %s | " (is_OK_aux l |> string_of_bool); + Printf.printf "is_bounded: %s\n" (is_bounded v2 loop |> string_of_bool); + (*is_nop v1 && *)is_branching v2 && is_OK_aux l && is_bounded v2 loop) + | _ -> false + +let print_loops loops = + List.iter (fun loop -> print_index (fst(loop)); + print_newline (); + List.iter print_index (snd(loop)); + print_newline (); + if is_pipelinable (snd(loop)) then print_string "PIPELINABLE !" else print_string "WASTE"; + print_newline (); + List.iter print_instruction (snd(loop)); + print_newline () + ) loops + +(* type resource = Reg of reg | Mem *) +module Sim = Map.Make (struct type t = resource let compare = compare end) + +let map_get key map = + try Some (Sim.find key map) + with + | Not_found -> None + +let rec to_res l = + match l with + | [] -> [] + | e :: l -> Reg e :: to_res l + +let resources_reads_of node = + match fst (G.V.label node) with + | Inop -> [] + | Iop (op, args, dst) -> to_res args + | Iload (chunk, mode, args, dst) -> Mem :: (to_res args) + | Istore (chunk, mode, args, src) -> Mem :: Reg src :: (to_res args) + | Icall (sign, id, args, dst) -> failwith "Resource read of call" + | Itailcall (sign, id, args) -> failwith "Resource read of tailcall" + (*| Ialloc (dst, size) -> [Mem] *) + | Icond (cond, args) -> to_res args + | Ireturn (res) -> failwith "Resource read of return" + +let resources_writes_of node = + match fst (G.V.label node) with + | Inop -> [] + | Iop (op, args, dst) -> [Reg dst] + | Iload (chunk, mode, args, dst) -> [Reg dst] + | Istore (chunk, mode, args, src) -> [Mem] + | Icall (sign, id, args, dst) -> failwith "Resource read of call" + | Itailcall (sign, id, args) -> failwith "Resource read of tailcall" + (*| Ialloc (dst, size) -> (Reg dst) :: [Mem]*) + | Icond (cond, args) -> [] + | Ireturn (res) -> failwith "Resource read of return" + +let build_intra_dependency_graph loop = + + let rec build_aux graph read write l = + match l with + | [] -> (graph,(read,write)) + | v :: l-> + let g = G.add_vertex graph v in + let reads = resources_reads_of v in + let writes = resources_writes_of v in + (* dependances RAW *) + let g = List.fold_right (fun r g -> + match map_get r write with + | Some n -> G.add_edge_e g (G.E.create n 1 v) + | None -> g + ) reads g in + (* dependances WAR *) + let g = List.fold_right (fun r g -> + match map_get r read with + | Some l -> List.fold_right (fun n g -> G.add_edge_e g (G.E.create n 2 v)) l g + | None -> g + ) writes g in + (* dependances WAW *) + let g = List.fold_right (fun r g -> + match map_get r write with + | Some n -> G.add_edge_e g (G.E.create n 3 v) + | None -> g + ) writes g in + let write = List.fold_right (fun r m -> Sim.add r v m) writes write in + let read_tmp = List.fold_right (fun r m -> + match map_get r read with + | Some al -> Sim.add r (v :: al) m + | None -> Sim.add r (v :: []) m + ) reads read + in + let read = List.fold_right (fun r m -> Sim.add r [] m) writes read_tmp in + + build_aux g read write l + in + + build_aux G.empty Sim.empty Sim.empty (List.tl loop) + +let build_inter_dependency_graph loop = + + let rec build_aux2 graph read write l = + match l with + | [] -> graph + | v :: l-> + let g = graph in + let reads = resources_reads_of v in + let writes = resources_writes_of v in + (* dependances RAW *) + let g = List.fold_right (fun r g -> + match map_get r write with + | Some n -> (* if n = v then g else *) G.add_edge_e g (G.E.create n 4 v) + | None -> g + ) reads g in + (* dependances WAR *) + let g = List.fold_right (fun r g -> + match map_get r read with + | Some l -> List.fold_right + (fun n g -> (* if n = v then g else *) G.add_edge_e g (G.E.create n 5 v)) l g + | None -> g + ) writes g in + (* dependances WAW *) + let g = List.fold_right (fun r g -> + match map_get r write with + | Some n -> (* if n = v then g else *) G.add_edge_e g (G.E.create n 6 v) + | None -> g + ) writes g in + let write = List.fold_right (fun r m -> Sim.remove r m) writes write in + let read = List.fold_right (fun r m -> Sim.remove r m) writes read in + + + build_aux2 g read write l + in + + let (g,(r,w)) = build_intra_dependency_graph loop in + build_aux2 g r w (List.tl loop) + +(* patch_graph prepare le graphe pour la boucle commencant au noeud entry + qui a une borne de boucle bound et pour un software pipelining + de au minimum min tours et de deroulement ur *) +(* this is rather technical so we give many comments *) + +(* let n1 = G.V.create (Iop ((Ointconst ur),to_coq_list [],r1),next_pc) in *) +(* let next_pc = next_pc + 1 in *) +(* let n2 = G.V.create (Iop (Odiv,to_coq_list [bound;r1],r2),next_pc) in *) +(* let next_pc = next_pc + 1 in *) +(* let n3 = G.V.create (Iop ((Omulimm ur),to_coq_list [r2],r3),next_pc) in *) +(* let next_pc = next_pc + 1 in *) +(* let n4 = G.V.create (Iop (Osub,to_coq_list [bound;r3],r4),next_pc) in *) +(* let next_pc = next_pc + 1 in *) +(* let n5 = G.V.create (Iop (Omove,to_coq_list [r3],bound),next_pc) in (\* retouchee, [r3],bound *\) *) + + +let patch_graph graph entry lastone loop bound min ur r1 r2 r3 r4 next_pc prolog epilog ramp_up ramp_down = + + (* 1. Break the edge that enters the loop, except for the backedge *) + let preds_e = G.pred_e graph entry in + let wannabes = List.map G.E.src preds_e in + let wannabes = List.filter (fun e -> not (e = lastone)) wannabes in + let graph = List.fold_right (fun e g -> G.remove_edge_e g e) preds_e graph in + let graph = G.add_edge graph lastone entry in + + (* 2. Add the test for minimal iterations and link it*) + + let cond = G.V.create (Icond ((Ccompimm (Integers.Cle,min)),to_coq_list [bound]), next_pc) in + let graph = G.add_vertex graph cond in + let next_pc = next_pc + 1 in + + (* 3. Link its predecessors and successors *) + (* It is false in case there is a condition that points to the entry: + inthis case, the edge should not be labeled with 0 !*) + + let graph = List.fold_right (fun n g -> G.add_edge g n cond) wannabes graph in + let graph = G.add_edge_e graph (G.E.create cond 1 entry) in + + + (* 4. Add the div and modulo code, link it *) + let n1 = G.V.create (Iop ((Ointconst ur),to_coq_list [],r1),next_pc) in + let next_pc = next_pc + 1 in + let n2 = G.V.create (Iop (Odiv,to_coq_list [bound;r1],r2),next_pc) in + let next_pc = next_pc + 1 in + let n3 = G.V.create (Iop (Omove,to_coq_list [bound],r4),next_pc) in + let next_pc = next_pc + 1 in + let n4 = G.V.create (Iop ((Omulimm ur ),to_coq_list [r2],r3),next_pc) in + let next_pc = next_pc + 1 in + let n5 = G.V.create (Iop ((Olea (Aindexed (Z.of_sint (-1)))),to_coq_list [r3],bound),next_pc) in (* retouchee, [r3],bound *) + let next_pc = next_pc + 1 in + let graph = G.add_vertex graph n1 in + let graph = G.add_vertex graph n2 in + let graph = G.add_vertex graph n3 in + let graph = G.add_vertex graph n4 in + let graph = G.add_vertex graph n5 in + let graph = G.add_edge_e graph (G.E.create cond 2 n1) in + let graph = G.add_edge graph n1 n2 in + let graph = G.add_edge graph n2 n3 in + let graph = G.add_edge graph n3 n4 in + let graph = G.add_edge graph n4 n5 in + + (* 5. Fabriquer la pipelined loop et la linker, sans la condition d entree *) + + let (graph,next_pc,l) = List.fold_right (fun e (g,npc,l) -> + let n = G.V.create (e,npc) in + (G.add_vertex g n, npc+1, n :: l) + ) loop (graph,next_pc,[]) in + + let pipe_cond = List.hd l in + let pipeline = List.tl l in + + let rec link l graph node = + match l with + | n1 :: n2 :: l -> link (n2 :: l) (G.add_edge graph n1 n2) node + | n1 :: [] -> G.add_edge graph n1 node + | _ -> graph + in + + let graph = link pipeline graph pipe_cond in + + (* link de l entree de la boucle *) + + let (graph,next_pc,prolog) = List.fold_right (fun e (g,npc,l) -> + let n = G.V.create (e,npc) in + (G.add_vertex g n, npc+1, n :: l) + ) prolog (graph,next_pc,[]) in + + let (graph,next_pc,epilog) = List.fold_right (fun e (g,npc,l) -> + let n = G.V.create (e,npc) in + (G.add_vertex g n, npc+1, n :: l) + ) epilog (graph,next_pc,[]) in + + (* 6. Creation du reste et branchement et la condition *) + let n6 = G.V.create (Iop (Omove,to_coq_list [r4],bound),next_pc) in (* Iop (Omove,to_coq_list [r4],bound) *) + let next_pc = next_pc + 1 in + + (* 7. Creation du ramp up *) + let ramp_up = List.map (fun (a,b) -> Iop (Omove, [b], a)) ramp_up in + let (graph,next_pc,ramp_up) = List.fold_right (fun e (g,npc,l) -> + let n = G.V.create (e,npc) in + (G.add_vertex g n, npc+1, n :: l) + ) ramp_up (graph,next_pc,[]) in + + let next_pc = next_pc + 1 in + + let ramp_down = List.map (fun (a,b) -> Iop (Omove,[b],a)) ramp_down in + let (graph,next_pc,ramp_down) = List.fold_right (fun e (g,npc,l) -> + let n = G.V.create (e,npc) in + (G.add_vertex g n, npc+1, n :: l) + ) ramp_down (graph,next_pc,[]) in + + (* let next_pc = next_pc + 1 in *) + + (* Creation des proloque et epilogue *) + + let graph = link prolog graph pipe_cond in + let graph = link ramp_up graph (List.hd prolog) in + let graph = link epilog graph (List.hd ramp_down) in + let graph = link ramp_down graph n6 in + + let graph = G.add_edge graph n5 (List.hd ramp_up) in + let graph = G.add_edge_e graph (G.E.create pipe_cond 1 (List.hd epilog)) in + let graph = G.add_edge_e graph (G.E.create pipe_cond 2 (List.hd pipeline)) in + + (* 8. Retour sur la boucle classique *) + let graph = G.add_edge graph n6 entry in + + graph + +let regs_of_node node = + match G.V.label node with + | (Inop,n) -> [] + | (Iop (op, args, dst),n) -> dst :: (to_caml_list args) + | (Iload (chunk, mode, args, dst),n) -> dst :: (to_caml_list args) + | (Istore (chunk, mode, args, src),n) -> src :: (to_caml_list args) + | (Icall (sign, id, args, dst),n) -> dst :: (to_caml_list args) + | (Itailcall (sign, id, args),n) -> (to_caml_list args) + (*| (Ialloc (dst, size),n) -> [dst]*) + | (Icond (cond, args),n) -> (to_caml_list args) + | (Ireturn (res),n) -> match res with Some res -> [res] | _ -> [] + +let max_reg_of_graph graph params = + Printf.fprintf SPDebug.dc "Calcul du registre de depart.\n"; + let regs = G.fold_vertex (fun node regs -> + (regs_of_node node) @ regs + ) graph [] in + let regs = regs @ params in + let max_reg = List.fold_right (fun reg max -> + Printf.fprintf SPDebug.dc "%i " (P.to_int reg); + if Int32.compare (P.to_int32 reg) max > 0 + then (P.to_int32 reg) + else max + ) regs Int32.zero in + + Printf.fprintf SPDebug.dc "MAX REG = %i\n" (Int32.to_int max_reg); + BinPos.Pos.succ (P.of_int32 max_reg) + +let get_bound node loop = + match G.V.label node with + | (Icond (cond, args),n) -> + let args = to_caml_list args in + begin + match args with + | [] -> failwith "get_bound: condition sans variables" + | r :: [] -> if is_variant r loop then failwith "Pas de borne dans la boucle" else r (* Modified false to true condition. *) + | r1 :: r2 :: [] -> + begin + match is_variant r1 loop, is_variant r2 loop with + | true, true -> failwith "Pas de borne dans la boucle " + | false, true -> r1 + | true, false -> r2 + | false, false -> failwith "deux bornes possibles dans la boucle" + end + | _ -> failwith "get_bound: condition avec nombre de variables superieur a 2" + end + | _ -> failwith "get_bound: the node I was given is not a condition\n" + +let get_nextpc graph = + (G.fold_vertex (fun node max -> + if (snd (G.V.label node)) > max + then (snd (G.V.label node)) + else max + ) graph 0) + 1 + +let substitute_pipeline graph loop steady_state prolog epilog min unrolling ru rd params = + let n1 = max_reg_of_graph graph params in + let n2 = (BinPos.Pos.succ n1) in + let n3 = (BinPos.Pos.succ n2) in + let n4 = (BinPos.Pos.succ n3) in + let way_in = (List.hd loop) in + let way_out = (List.hd (List.rev loop)) in + let bound = (get_bound way_out loop) in + let min = Z.of_sint min in + let unrolling = Z.of_sint unrolling in + let next_pc = get_nextpc graph in + patch_graph graph way_in way_out steady_state bound min unrolling n1 n2 n3 n4 next_pc prolog epilog ru rd + +let get_loops cfg = + let domi = dominator_tree cfg in + let loops = detect_loops cfg.graph domi in + print_loops loops; + let loops = List.filter (fun loop -> is_pipelinable (snd (loop))) loops in + loops + +type pipeline = {steady_state : G.V.t list; prolog : G.V.t list; epilog : G.V.t list; + min : int; unrolling : int; ramp_up : (reg * reg) list; ramp_down : (reg * reg) list} + +let delete_indexes l = List.map (fun e -> fst (G.V.label e) ) l + +type reg = Registers.reg + +let fresh = ref BinNums.Coq_xH + +let distance e = + match G.E.label e with + | 1 | 2 | 3 -> 0 + | _ -> 1 + +type et = IntraRAW | IntraWAW | IntraWAR | InterRAW | InterWAW | InterWAR + +let edge_type e = + match G.E.label e with + | 1 -> IntraRAW + | 2 -> IntraWAR + | 3 -> IntraWAW + | 4 -> InterRAW + | 5 -> InterWAR + | 6 -> InterWAW + | _ -> failwith "Unknown edge type" + +let latency n = (* A raffiner *) + match fst (G.V.label n) with + | Iop (op,args, dst) -> + begin + match op with + | Omove -> 1 + (*| Oaddimm _ -> 1*) + (*| Oadd -> 2*) + | Omul -> 4 + | Odiv -> 30 + | Omulimm _ -> 4 + | _ -> 2 + end + | Iload _ -> 1 + (* | Ialloc _ -> 20*) + | _ -> 1 + +let reforge_writes inst r = + G.V.create ((match fst (G.V.label inst) with + | Inop -> Inop + | Iop (op, args, dst) -> Iop (op, args, r) + | Iload (chunk, mode, args, dst) -> Iload (chunk, mode, args, r) + | Istore (chunk, mode, args, src) -> Istore (chunk, mode, args, src) + | Icall (sign, id, args, dst) -> failwith "reforge_writes: call" + | Itailcall (sign, id, args) -> failwith "reforge_writes: tailcall" + (* | Ialloc (dst, size) -> Ialloc (r, size)*) + | Icond (cond, args) -> Icond (cond, args) + | Ireturn (res) -> failwith "reforge_writes: return") + , snd (G.V.label inst)) + +let rec reforge_args args oldr newr = + match args with + | [] -> [] + | e :: l -> (if e = oldr then newr else e) :: (reforge_args l oldr newr) + +let rec mem_args args r = + match args with + | [] -> false + | e :: l -> if e = r then true else mem_args l r + +let check_read_exists inst r = + match fst (G.V.label inst) with + | Inop -> false + | Iop (op, args, dst) -> mem_args args r + | Iload (chunk, mode, args, dst) -> mem_args args r + | Istore (chunk, mode, args, src) -> src = r || mem_args args r + | Icall (sign, id, args, dst) -> mem_args args r + | Itailcall (sign, id, args) -> false + (*| Ialloc (dst, size) -> false*) + | Icond (cond, args) -> mem_args args r + | Ireturn (res) -> false + +let reforge_reads inst oldr newr = + assert (check_read_exists inst oldr); + G.V.create ((match fst (G.V.label inst) with + | Inop -> Inop + | Iop (op, args, dst) -> Iop (op, reforge_args args oldr newr, dst) + | Iload (chunk, mode, args, dst) -> Iload (chunk, mode, reforge_args args oldr newr, dst) + | Istore (chunk, mode, args, src) -> Istore (chunk, mode, reforge_args args oldr newr , if src = oldr then newr else src) + | Icall (sign, id, args, dst) -> failwith "reforge_reads: call" + | Itailcall (sign, id, args) -> failwith "reforge_reads: tailcall" + (*| Ialloc (dst, size) -> Ialloc (dst, size)*) + | Icond (cond, args) -> Icond (cond, reforge_args args oldr newr) + | Ireturn (res) -> failwith "reforge_reads: return") + , snd (G.V.label inst)) + +let get_succs_raw ddg node = + let succs = G.succ_e ddg node in + let succs = List.filter (fun succ -> + match G.E.label succ with + | 1 | 4 -> true + | _ -> false + ) succs in + List.map (fun e -> G.E.dst e) succs + +let written inst = + match fst (G.V.label inst) with + | Inop -> None + | Iop (op, args, dst) -> Some dst + | Iload (chunk, mode, args, dst) -> Some dst + | Istore (chunk, mode, args, src) -> None + | Icall (sign, id, args, dst) -> failwith "written: call" + | Itailcall (sign, id, args) -> failwith "written: tailcall" + (*| Ialloc (dst, size) -> Some dst*) + | Icond (cond, args) -> None + | Ireturn (res) -> failwith "written: return" + +let fresh_regs n = + let t = Array.make n (BinNums.Coq_xH) in + for i = 0 to (n - 1) do + Array.set t i (!fresh); + fresh := BinPos.Pos.succ !fresh + done; + t + +let print_reg r = Printf.fprintf SPDebug.dc "%i " (P.to_int r) + +let is_cond node = + match fst (G.V.label node) with + | Icond _ -> true + | _ -> false + + +(*******************************************) + +let watch_regs l = List.fold_right (fun (a,b) l -> + if List.mem a l then l else a :: l + ) l [] + +let make_moves = List.map (fun (a,b) -> Iop (Omove,[b],a)) + +let rec repeat l n = + match n with + | 0 -> [] + | n -> l @ repeat l (n-1) + +let fv = ref 0 + +let apply_pipeliner f p ?(debug=false) = + Printf.fprintf SPDebug.dc "******************** NEW FUNCTION ***********************\n"; + let cfg = convert f in + incr fv; + if debug then display cfg.graph ("input" ^ (string_of_int !fv)); + let loops = get_loops cfg in + Printf.fprintf SPDebug.dc "Loops: %d\n" (List.length loops); + let ddgs = List.map (fun (qqch,loop) -> (loop,build_inter_dependency_graph loop)) loops in + + let lv = ref 0 in + + let graph = List.fold_right (fun (loop,ddg) graph -> + Printf.fprintf SPDebug.dc "__________________ NEW LOOP ____________________\n"; + Printf.printf "Pipelinable loop: "; + incr lv; + fresh := (BinPos.Pos.succ + (BinPos.Pos.succ + (BinPos.Pos.succ + (BinPos.Pos.succ + (BinPos.Pos.succ + (max_reg_of_graph graph (to_caml_list f.fn_params) + )))))); + Printf.fprintf SPDebug.dc "FRESH = %i \n" + (P.to_int !fresh); + match p ddg with + | Some pipe -> + Printf.printf "Rock On ! Min = %i - Unroll = %i\n" pipe.min pipe.unrolling; + let p = (make_moves pipe.ramp_up) @ (delete_indexes pipe.prolog) in + let e = (delete_indexes pipe.epilog) @ (make_moves pipe.ramp_down) in + let b = delete_indexes (List.tl (List.rev loop)) in + let bt = (List.tl (delete_indexes pipe.steady_state)) in + let cond1 = fst (G.V.label (List.hd (List.rev loop))) in + let cond2 = (List.hd (delete_indexes pipe.steady_state)) in + + let bu = symbolic_evaluation (repeat b (pipe.min + 1)) in + let pe = symbolic_evaluation (p @ e) in + let bte = symbolic_evaluation (bt @ e) in + let ebu = symbolic_evaluation (e @ repeat b pipe.unrolling) in + let regs = watch_regs pipe.ramp_down in + let c1 = symbolic_condition cond1 (repeat b pipe.unrolling) in + let d1 = symbolic_condition cond1 (repeat b (pipe.min + 1)) in + (*let c2 = symbolic_condition cond2 p in + let d2 = symbolic_condition cond2 ((make_moves pipe.ramp_up) @ bt) in*) + + + + let sbt = symbolic_evaluation (bt) in + let sep = symbolic_evaluation (e @ repeat b (pipe.unrolling - (pipe.min + 1)) @ p) in (* er @ pr *) + + Printf.printf "Initialisation : %s \n" + (if symbolic_equivalence bu pe regs then "OK" else "FAIL"); + Printf.printf "Etat stable : %s \n" + (if symbolic_equivalence bte ebu regs then "OK" else "FAIL"); + Printf.printf "Egalite fondamentale : %s \n" + (if symbolic_equivalence sbt sep (watch_regs pipe.ramp_up) then "OK" else "FAIL"); + (* Printf.printf "Condition initiale : %s \n" + (if c1 = c2 then "OK" else "FAIL"); + Printf.printf "Condition stable : %s \n" + (if d1 = d2 then "OK" else "FAIL"); + + + Printf.fprintf SPDebug.dc "pbte\n";*) + List.iter (fun e -> + Printf.fprintf SPDebug.dc "%s\n" + (string_of_node (G.V.create (e,0))) + ) (p @ bt @ e); + Printf.fprintf SPDebug.dc "bu\n"; + List.iter (fun e -> Printf.fprintf SPDebug.dc "%s\n" + (string_of_node (G.V.create (e,0))) + ) (repeat b (pipe.unrolling + pipe.min)); + + + + if debug then + display_st ("pbte"^ (string_of_int !fv) ^ (string_of_int !lv)) (p @ bt @ e) (watch_regs pipe.ramp_down); + if debug then + display_st ("bu"^ (string_of_int !fv) ^ (string_of_int !lv)) (repeat b (pipe.min + pipe.unrolling)) (watch_regs pipe.ramp_down); + + if debug then display_st ("bt"^ (string_of_int !fv) ^ (string_of_int !lv)) (bt) (watch_regs pipe.ramp_up); + if debug then display_st ("ep"^ (string_of_int !fv) ^ (string_of_int !lv)) (e @ repeat b (pipe.unrolling - (pipe.min + 1)) @ p) (watch_regs pipe.ramp_up); + + substitute_pipeline graph loop + (delete_indexes pipe.steady_state) (delete_indexes pipe.prolog) + (delete_indexes pipe.epilog) (pipe.min + pipe.unrolling) + pipe.unrolling pipe.ramp_up + pipe.ramp_down + (to_caml_list f.fn_params) + | None -> Printf.printf "Damn It ! \n"; graph + ) ddgs cfg.graph in + + if debug then display graph ("output"^ (string_of_int !fv)); + let tg = convert_back graph in + + let tg_to_type = {fn_sig = f.fn_sig; + fn_params = f.fn_params; + fn_stacksize = f.fn_stacksize; + fn_code = tg; + fn_entrypoint = f.fn_entrypoint; + (*fn_nextpc = P.of_int ((get_nextpc (graph)))*) + } in + (*SPTyping.type_function tg_to_type;*) + + tg_to_type diff --git a/src/pipelining/SPBasic.mli b/src/pipelining/SPBasic.mli new file mode 100644 index 0000000..f16e524 --- /dev/null +++ b/src/pipelining/SPBasic.mli @@ -0,0 +1,57 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open Graph.Pack.Digraph + +(* DATA DEPENDENCY GRAPHS *) +module G : Graph.Sig.P + +(* We abstract the register type to make sure that the user won't mess up *) +type reg + +(* The scheduling should return a value of type pipeline *) +type pipeline = { + steady_state : G.V.t list; + prolog : G.V.t list; + epilog : G.V.t list; + min : int; + unrolling : int; + ramp_up : (reg * reg) list; + ramp_down : (reg * reg) list +} + +(* Operations on ddg *) + +val display : G.t -> string -> unit +val apply_pipeliner : RTL.coq_function -> (G.t -> pipeline option) -> ?debug:bool -> RTL.coq_function +val get_succs_raw : G.t -> G.V.t -> G.V.t list + +(* Operations on instructions, the nodes of the data dependency graph *) + +val string_of_node : G.V.t -> string +val latency : G.V.t -> int +val reforge_reads : G.V.t -> reg -> reg -> G.V.t +val reforge_writes : G.V.t -> reg -> G.V.t +val written : G.V.t -> reg option +val is_cond : G.V.t -> bool + +(* Operations on dependencies, the edges of the data dependency graph *) + +type et = IntraRAW | IntraWAW | IntraWAR | InterRAW | InterWAW | InterWAR + +val edge_type : G.E.t -> et +val distance : G.E.t -> int + +(* Getting fresh registers, int is the number of required registers *) + +val fresh_regs : int -> reg array +val print_reg : reg -> unit diff --git a/src/pipelining/SPDebug.ml b/src/pipelining/SPDebug.ml new file mode 100644 index 0000000..34d1c0c --- /dev/null +++ b/src/pipelining/SPDebug.ml @@ -0,0 +1,21 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open Unix + +let tm = localtime (time ());; +let name = "debug/" ^ (string_of_int tm.tm_year) ^ "-" ^(string_of_int tm.tm_mon) ^ "-" ^(string_of_int tm.tm_mday) ^ + "-" ^(string_of_int tm.tm_hour) ^"-" ^(string_of_int tm.tm_min) ^ "-" ^(string_of_int tm.tm_sec) ^ "/";; +mkdir name 0o777;; +Printf.printf "Debug informations: %s \n" name ;; +let dc = open_out (name ^ "debug.log");; +let () = at_exit(fun () -> close_out dc);; diff --git a/src/pipelining/SPIMS.ml b/src/pipelining/SPIMS.ml new file mode 100644 index 0000000..0e19dec --- /dev/null +++ b/src/pipelining/SPIMS.ml @@ -0,0 +1,189 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open Graph.Pack.Digraph +open SPBasic + +module NI = Map.Make (struct type t = G.V.t let compare = compare end) + +let find key map def = + try NI.find key map + with + | Not_found -> def + +let unpack v = + match v with + | Some v -> v + | None -> failwith "unpack abusif" + +let dep_latency edge = + match edge_type edge with + | IntraRAW | InterRAW -> latency (G.E.src edge) + | _ -> 1 + +let estart ddg schedule node ii = + let preds = G.pred_e ddg node in + let starts = List.map (fun edge -> + match find (G.E.src edge) schedule None with + | Some t -> + let start = t + dep_latency edge - ii * distance edge in + (*Printf.printf "start : %i \n" start;*) + if start < 0 then 0 else start + | None -> 0 + ) preds in + List.fold_left (fun max e -> if max > e then max else e) 0 starts + +let resource_conflict time mrt ii = + match Array.get mrt (time mod ii) with + | None -> false + | Some v -> true + +let rec scan_time time maxtime mrt ii = + if time <= maxtime + then + begin + if resource_conflict time mrt ii + then scan_time (time + 1) maxtime mrt ii + else Some time + end + else None + +let finished ddg schedule = + let unscheduled = G.fold_vertex (fun node l -> + match find node schedule None with + | Some v -> l + | None -> node :: l + ) ddg [] in + (* Printf.printf "R %i R \n" (List.length unscheduled); *) + if List.length unscheduled = 0 then true else false + +let bad_successors ddg schedule node ii = + let succs = G.succ_e ddg node in (* Le succs_ddg initial *) + (* let reftime = NI.find node schedule in *) + (* let succs_sched = NI.fold (fun node time succs -> *) + (* if time > reftime then node :: succs else succs *) + (* ) schedule [] in *) + (* let succs = List.filter (fun edge -> List.mem (G.E.dst edge) succs_sched) succs_ddg in*) + List.fold_right (fun edge bad -> + match find (G.E.dst edge) schedule None with + | Some t -> + if unpack (NI.find node schedule) + dep_latency edge - ii * distance edge > t + then (G.E.dst edge) :: bad + else bad + | None -> bad + ) succs [] + +let get_condition ddg = + let cond = G.fold_vertex (fun node cond -> + if is_cond node then Some node else cond + ) ddg None in + match cond with + | Some cond -> cond + | None -> failwith "The loop does not contain a condition. Aborting\n" + +(* Perform iterative modulo scheduling, using a heuristic for the next instruction to schedule + * [heightr], the data dependency graph to schedule [ddg], the minimum II [min_ii] and the maximum + * II [max_interval]. + *) +let modulo_schedule heightr ddg min_ii max_interval = + + let ii = ref (min_ii - 1) in + let notfound = ref true in + let sched = ref NI.empty in + + let cond = get_condition ddg in + + while (!ii < max_interval && !notfound) do + (* Printf.printf "."; flush stdout; *) + ii := !ii + 1; + (* Printf.printf "NOUVEAU II %i \n" !ii; *) + let budget = ref (G.nb_vertex ddg * 10) in + let lasttime = ref NI.empty in + (* Create the map with schedules, and add the schedule for the condition. This should go at the + * end, but in this case is set to the start. *) + let schedtime = ref (NI.add cond (Some 0) NI.empty) in + (* Create an array that is as large as the current II, which will determine where each + * instruction will be placed. *) + let mrt = Array.make !ii None in + (* Set the condition to be the initial instruction at time 0. *) + Array.set mrt 0 (Some cond); + + while (!budget > 0 && not (finished ddg !schedtime)) do (* Pretty inefficient *) + budget := !budget - 1; + (* Get next instruction to schedule. *) + let h = heightr ddg !schedtime in + let mintime = estart ddg !schedtime h !ii in + (* Printf.printf "tmin (%s) = %i \n" (string_of_node h) mintime; *) + let maxtime = mintime + !ii -1 in + let time = + match scan_time mintime maxtime mrt !ii with + | Some t -> t + | None -> (*Printf.printf "backtrack" ; *) + if mintime = 0 then 1 + find h !lasttime 0 + else max mintime (1 + find h !lasttime 0) + in + (* Printf.printf "Chosen time for %s : %i \n" (string_of_node h) time; *) + schedtime := NI.add h (Some time) !schedtime; + lasttime := NI.add h time !lasttime; + + let killed = bad_successors ddg !schedtime h !ii in + List.iter (fun n -> (* Printf.printf "Killing %s" (string_of_node n) ; *)schedtime := NI.add n None !schedtime) killed; + + begin + match Array.get mrt (time mod !ii) with + | None -> Array.set mrt (time mod !ii) (Some h) + | Some n -> + begin + (*Printf.printf "Deleting : %s \n" (string_of_node n); *) + (* Printf.printf "."; *) + schedtime := NI.add n None !schedtime; + Array.set mrt (time mod !ii) (Some h) + end + end; + (* if finished ddg !schedtime then Printf.printf "Fini ! \n" *) + + done; + + let success = G.fold_vertex (fun node b -> + b && + match find node !schedtime None with + | Some _ -> true + | None -> false + ) ddg true in + + if success then (notfound := false; sched := !schedtime); + + done; + + if (not !notfound) + then (!sched,!ii) + else failwith "IMS failure" + +(* Take the number of vertices as the minimum resource-constrained II. However, the II might + actually be less than that in some cases, so this should be adjusted accordingly. *) +let res_m_ii ddg = + G.nb_vertex ddg + +let pipeliner ddg heightr = + let mii = res_m_ii ddg in + Printf.fprintf SPDebug.dc "MII: %i\n" mii; + let (schedule,ii) = modulo_schedule heightr ddg mii (10 * mii) in + (NI.fold (fun n v map -> + match v with + | Some v -> NI.add n v map + | None -> failwith "pipeliner: schedule unfinished" + ) schedule NI.empty,ii) + +let print_schedule sched = + NI.iter (fun node time -> + Printf.fprintf SPDebug.dc "%s |---> %i \n" (string_of_node node) time + ) sched diff --git a/src/pipelining/SPIMS.mli b/src/pipelining/SPIMS.mli new file mode 100644 index 0000000..7c1d9a7 --- /dev/null +++ b/src/pipelining/SPIMS.mli @@ -0,0 +1,22 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open Graph.Pack.Digraph +open SPBasic + +module NI : Map.S with type key = SPBasic.G.V.t + +(* piepeliner takes a data dependency graph and returns a schedule with an initiation interval + fails if cannot find any schedule *) +val pipeliner : G.t -> (G.t -> int option NI.t -> G.V.t) -> int NI.t * int + +val print_schedule : int NI.t -> unit diff --git a/src/pipelining/SPMVE.ml b/src/pipelining/SPMVE.ml new file mode 100644 index 0000000..28381c1 --- /dev/null +++ b/src/pipelining/SPMVE.ml @@ -0,0 +1,299 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open SPBasic +open SPIMS + +module Mult = Map.Make (struct type t = reg let compare = compare end) + +let size_of_map1 m = + NI.fold (fun key v size -> size + 1) m 0 + +let size_of_map2 m = + Mult.fold (fun key v size -> size + 1) m 0 + +let sched_max_time sched = + NI.fold (fun node time max -> + if time > max then time else max + ) sched 0 + +let print_table t s = + Printf.fprintf SPDebug.dc "%s : \n" s; + let string_of_node_ev node = + match node with + | Some node -> string_of_node node + | None -> "_" + in + Array.iteri (fun i node -> Printf.fprintf SPDebug.dc "%i :: %s \n" i (string_of_node_ev node)) t + +let durations ddg sched ii = + + G.fold_vertex (fun node mult -> + match written node with + | None -> mult + | Some r -> + let raw_succs = get_succs_raw ddg node in + let durations = List.map (fun succ -> + let d = NI.find succ sched - NI.find node sched in + if d >= 0 then d + else ((sched_max_time sched - NI.find node sched) + NI.find succ sched) + ) raw_succs in + let duration = List.fold_left (fun max e -> if max > e then max else e) 0 durations in + Mult.add r ((duration / ii) + 1) mult (* cette division est surement fdausse*) + ) ddg Mult.empty + +let minimizer qi ur = + + let rec fill n = + if n <= ur then n :: fill (n + 1) else [] + in + + let l = fill qi in + let l = List.map (fun e -> (e,ur mod e)) l in + let l = List.filter (fun e -> snd e = 0) l in + let l = List.map fst l in + List.fold_left (fun min e -> if min < e then min else e) ur l + +let multiplicity ddg sched ii = + let qs = durations ddg sched ii in + (* Printf.printf "Quantite de variables necessaires : \n"; *) + (* Mult.iter (fun key mu -> print_reg key ; Printf.printf " |-> %i\n" mu) qs; *) + let unroll = Mult.fold (fun reg mult max -> + if mult > max then mult else max + ) qs 0 in + let mult = Mult.fold (fun reg mult mult -> + Mult.add reg (minimizer (Mult.find reg qs) unroll) mult + ) qs Mult.empty + in + (mult,unroll) + +let mve_kernel t ddg sched ii mult unroll = + + let regs = Array.make ii (fresh_regs 0) in + for i = 0 to (ii - 1) do + let fregs = fresh_regs unroll in + Array.iter print_reg fregs; + Printf.fprintf SPDebug.dc "\n"; + Array.set regs i fregs + done; + + let used_regs = ref [] in + + let index r i = Mult.find r mult - ( ((i / ii) + 1) mod Mult.find r mult) in + (* let pos i node inst = *) + (* let separation = *) + (* let b= NI.find node sched - NI.find inst sched in *) + (* if b >= 0 then b *) + (* else ((sched_max_time sched - NI.find inst sched) + NI.find node sched) *) + (* in *) + (* (i + separation) mod (ii * unroll) in *) + + let new_t = Array.copy t in + + for i = 0 to (Array.length t - 1) do + (* print_table new_t "Nouvelle table"; *) + match t.(i),new_t.(i) with + | Some insti, Some insti_mod -> + begin + match written insti with + | None -> () + | Some r -> + begin + let new_reg = + if Mult.find r mult = 0 then r + else regs.(i mod ii).(index r i - 1) in + if (not (List.mem (r,new_reg) !used_regs)) then used_regs := (r,new_reg) :: !used_regs; + let inst = reforge_writes insti_mod new_reg in + Printf.fprintf SPDebug.dc "Reecriture : %s --> %s \n" (string_of_node insti) (string_of_node inst); + Array.set new_t i (Some inst); + let succs = get_succs_raw ddg insti in + List.iter (fun node -> + + (* let lifetime = *) + (* let d = NI.find node sched - NI.find insti sched in *) + (* if d >= 0 then d *) + (* else ((sched_max_time sched - NI.find insti sched) + 1 + NI.find node sched) *) + (* in *) + + + (* ___Version 1___ *) + (* let lifetime = lifetime / ii in *) + (* let schedtime = *) + (* ((NI.find node sched) mod ii) (\* Position dans le premier bloc *\) *) + (* + (ii * (i / ii)) (\* Commencement du bloc ou on travail *\) *) + (* + (ii * lifetime ) (\* Decalage (Mult.find r mult - 1) *\) *) + (* + ((if (NI.find node sched mod ii) < (NI.find insti sched mod ii) then 0 else 0) * ii) *) + (* in *) + + (* Printf.printf "seed = %i - bloc = %i - slide = %i - corr = %i - id = %i \n" *) + (* ((NI.find node sched) mod ii) *) + (* (ii * (i / ii)) (ii * lifetime) *) + (* ((if (NI.find node sched mod ii) < (NI.find insti sched mod ii) then 1 else 0 ) * ii) *) + (* id; *) + (* Printf.printf "Successeur a traiter : %s ooo %i ooo\n" (string_of_node node) (Mult.find r mult); *) + + (* ___Version 2___ *) + let schedtime = + if (NI.find node sched > NI.find insti sched) + then i + (NI.find node sched - NI.find insti sched) + else i - NI.find insti sched + ii + NI.find node sched + (* let delta = NI.find insti sched - NI.find node sched in *) + (* (i - delta) + (((delta / ii) + 1) * ii) (\* (i - delta) + ii*\) *) + in + + (* then *) + + Printf.fprintf SPDebug.dc "i = %i - def = %i - use = %i - time = %i \n" + i (NI.find insti sched) (NI.find node sched) schedtime; + + + (* let id = pos i node insti in *) + let id = schedtime mod (unroll * ii) (* le tout modulo la tabl *) in + let id = (id + (unroll * ii)) mod (unroll * ii) in + Printf.fprintf SPDebug.dc "Positions to treat : %i \n" id; + let insttt = match new_t.(id) with + | Some inst -> inst + | None -> failwith "There should be an instruction" + in + let inst = reforge_reads insttt r new_reg in + Array.set new_t id (Some inst) + ) succs + end + end + | None, _ -> () + | _, None -> failwith "MVE : qqch mal compris" + done; + new_t,!used_regs + +let crunch_and_unroll sched ii ur = + let steady_s = Array.make ii None in + NI.iter (fun inst time -> + Array.set steady_s (time mod ii) (Some inst) + ) sched; + (* Printf.printf "Etat stable : \n"; *) + (* let string_of_node_ev node = *) + (* match node with *) + (* | Some node -> string_of_node node *) + (* | None -> "_" *) + (* in *) + (* Array.iteri (fun i node -> Printf.printf "%i :: %s \n" i (string_of_node_ev node)) steady_s; *) + let steady = Array.make (ii * ur) None in + for i = 0 to (ur - 1) do + for time = 0 to (ii - 1) do + Array.set steady (time + i * ii) (steady_s.(time)) + done + done; + steady + +let compute_iteration_table sched ii = + let t = Array.make ii None in + NI.iter (fun node time -> + Array.set t (NI.find node sched mod ii) (Some ((NI.find node sched / ii) + 1)) + ) sched; + t + +let compute_prolog steady min ii unroll schedule it = + + let prolog = ref [] in + let prolog_piece = ref [] in + + for i = (min - 1) downto 0 do + + let index = ((ii * (unroll - (min - i)))) mod (unroll * ii) in + prolog_piece := []; + + for j = 0 to (ii - 1) do (* copie du sous tableau *) + (* Printf.printf "i : %i - j : %i - index : %i \n" i j index; *) + match steady.(index + j), it.(j) with + | Some inst , Some iter -> + if iter <= (i + 1) then prolog_piece := inst :: !prolog_piece; (* i + 1 au lieu de i *) + | None, _ -> () + | _, _ -> failwith "compute_prolog: quelquechose est mal compris" + done; + + prolog := List.rev (!prolog_piece) @ !prolog + done; + + !prolog + + +let compute_epilog steady min ii unroll schedule it = + + let epilog = ref [] in + + for i = 0 to (min - 1) do + let index = (i mod unroll) * ii in + for j = 0 to (ii - 1) do + match steady.(index + j), it.(j) with + | Some inst , Some iter -> + if iter > (i + 1) then epilog := inst :: !epilog; + | None, _ -> () + | _, _ -> failwith "compute_prolog: quelquechose est mal compris" + done; + done; + List.rev (!epilog) + +let entrance = List.map (fun (a,b) -> (b,a)) + +let way_out prolog epilog used_regs = + let l = List.rev (prolog @ epilog) in + + let rec way_out_rec l wo = + match l with + | [] -> wo + | i :: l -> + begin + match written i with + | Some r -> + let mov = List.find (fun (a,b) -> r = b) used_regs in + if List.mem mov wo + then way_out_rec l wo + else way_out_rec l (mov :: wo) + | None -> way_out_rec l wo + end + in + + way_out_rec l [] + +let mve ddg sched ii = + assert (size_of_map1 sched = G.nb_vertex ddg); + Printf.fprintf SPDebug.dc "L'intervalle d'initiation est de : %i \n" ii; + Printf.fprintf SPDebug.dc "L'ordonnancement est le suivant : \n"; + print_schedule sched; + let (mult,unroll) = multiplicity ddg sched ii in + let unroll = unroll in (* changer pour tester, default doit etre egal a unroll *) + Printf.fprintf SPDebug.dc "Table de multiplicite : \n"; + Mult.iter (fun key mu -> print_reg key ; Printf.fprintf SPDebug.dc " |-> %i\n" mu) mult; + Printf.fprintf SPDebug.dc "Taux de deroulement de : %i \n" unroll; + let steady_state = crunch_and_unroll sched ii unroll in + let (steady_state,used_regs) = mve_kernel steady_state ddg sched ii mult unroll in + print_table steady_state "Table finale"; + let min = ((sched_max_time sched) / ii) + 1 in + Printf.fprintf SPDebug.dc "min : %i \n" min; + let iteration_table = compute_iteration_table sched ii in + Printf.fprintf SPDebug.dc "Table d'iteration \n"; + Array.iteri (fun i elt -> + match elt with + | Some elt -> + Printf.fprintf SPDebug.dc "%i : %i\n" i elt + | None -> Printf.fprintf SPDebug.dc "%i : _ \n" i + ) iteration_table; + let prolog = compute_prolog steady_state min ii unroll sched iteration_table in + let prolog = List.filter (fun e -> not (is_cond e)) prolog in + let epilog = compute_epilog steady_state min ii unroll sched iteration_table in + let epilog = List.filter (fun e -> not (is_cond e)) epilog in + Printf.fprintf SPDebug.dc "Prologue: \n"; + List.iter (fun node -> Printf.fprintf SPDebug.dc "%s \n" (string_of_node node)) prolog; + Printf.fprintf SPDebug.dc "Epilogue: \n"; + List.iter (fun node -> Printf.fprintf SPDebug.dc "%s \n" (string_of_node node)) epilog; + let way_out = way_out prolog epilog used_regs in + (steady_state,prolog,epilog,min - 1,unroll,entrance used_regs, way_out) diff --git a/src/pipelining/SPMVE.mli b/src/pipelining/SPMVE.mli new file mode 100644 index 0000000..2da367d --- /dev/null +++ b/src/pipelining/SPMVE.mli @@ -0,0 +1,17 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open SPBasic +open SPIMS + +val mve : G.t -> int NI.t -> int -> + (G.V.t option) array * G.V.t list * G.V.t list * int * int * (reg * reg) list * (reg * reg) list diff --git a/src/pipelining/SPSymbolic_evaluation.ml b/src/pipelining/SPSymbolic_evaluation.ml new file mode 100644 index 0000000..c99afc8 --- /dev/null +++ b/src/pipelining/SPSymbolic_evaluation.ml @@ -0,0 +1,226 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open Registers +open Op +open AST +open SPBase_types +open Camlcoq + +type symbolic_value = + | Sreg of reg + | Sop of operation * symbolic_value list + | Sload of memory_chunk * addressing * symbolic_value list * symbolic_mem + +and symbolic_mem = + | Smem + | Sstore of memory_chunk * addressing * symbolic_value list * symbolic_value * symbolic_mem + +module State = Map.Make (struct type t = reg let compare = compare end) + +module Cons = Set.Make (struct type t = symbolic_value let compare = compare end) + +type symbolic_state = symbolic_value State.t * Cons.t + +let initial_state = State.empty +let initial_mem = Smem +let initial_cons = Cons.empty + +exception Not_straight + +let find res st = + try State.find res st + with + | Not_found -> Sreg res + +let rec get_args st = function + | [] -> [] + | arg::args -> find arg st :: get_args st args + +let rec symbolic_evaluation st sm cs = function + | [] -> (st,sm,cs) + | Inop :: l -> symbolic_evaluation st sm cs l + + | Iop (Omove, [src], dst) :: l -> + symbolic_evaluation (State.add dst (find src st) st) sm cs l + + | Iop (op, args, dst) :: l -> + let sym_val = Sop (op,get_args st args) in + symbolic_evaluation (State.add dst sym_val st) sm (Cons.add sym_val cs) l + + | Iload (chunk, mode, args, dst) :: l -> + let sym_val = Sload (chunk, mode, get_args st args, sm) in + symbolic_evaluation (State.add dst sym_val st) sm (Cons.add sym_val cs) l + + | Istore (chunk, mode, args, src) :: l -> + let sym_mem = Sstore (chunk, mode, get_args st args, find src st, sm) in + symbolic_evaluation st sym_mem cs l + + | _ :: l -> raise Not_straight + +type osv = + | Oresource of resource + | Oop of operation + | Oload of memory_chunk * addressing + | Ostore of memory_chunk * addressing + +let string_of_osv = function + | Oresource (Reg r) -> Printf.sprintf "reg %i" (P.to_int r) + | Oresource Mem -> "mem" + | Oop op -> string_of_op op + | Oload (mc,addr) -> "load" + | Ostore (mc,addr) -> "store" + +type ident = int + +module S = Graph.Persistent.Digraph.Abstract + (struct type t = osv * ident end) + +let name_of_vertex v = + let (osv,id) = S.V.label v in + Printf.sprintf "%i" id + +let string_of_vertex v = + let (osv,_) = S.V.label v in + Printf.sprintf "%s" (string_of_osv osv) + +module DisplayTree = struct + include S + let vertex_name v = name_of_vertex v + let graph_attributes _ = [] + let default_vertex_attributes _ = [] + let vertex_attributes v = [`Label (string_of_vertex v)] + let default_edge_attributes _ = [] + let edge_attributes _ = [] + let get_subgraph _ = None +end +module DotTree = Graph.Graphviz.Dot(DisplayTree) + +let dot_output_ss g f = + let oc = open_out f in + DotTree.output_graph oc g; + close_out oc + +module Build = Graph.Builder.P (S) +module Op = Graph.Oper.Make (Build) + +let counter = ref 0 + +let rec convert_sv_rec sv graph = + incr counter; + match sv with + | Sreg res -> + let node = S.V.create (Oresource (Reg res), !counter) in + let graph = S.add_vertex graph node in + (graph,node) + + | Sop (op, svl) -> + let node = S.V.create (Oop op, !counter) in + let (graph, node_l) = List.fold_right (fun sv (graph,node_l) -> + let (graph,node) = convert_sv_rec sv graph in + graph, node :: node_l + ) svl (graph,[]) in + let graph = S.add_vertex graph node in + let graph = List.fold_right (fun n graph -> + S.add_edge graph node n + ) node_l graph in + (graph,node) + + + | Sload (mc,addr,svl,sm) -> + let node = S.V.create (Oload (mc, addr), !counter) in + let (graph, node_l) = List.fold_right (fun sv (graph,node_l) -> + let (graph,node) = convert_sv_rec sv graph in + graph, node :: node_l + ) svl (graph,[]) in + let (graph,node_m) = convert_sm_rec sm graph in + let graph = S.add_vertex graph node in + let graph = List.fold_right (fun n graph -> + S.add_edge graph node n + ) node_l graph in + let graph = S.add_edge graph node node_m in + (graph,node) + +and convert_sm_rec sm graph = + incr counter; + match sm with + | Smem -> + let node = S.V.create (Oresource Mem, !counter) in + let graph = S.add_vertex graph node in + (graph,node) + + | Sstore (mc,addr,svl,sv,sm) -> + let node = S.V.create (Ostore (mc, addr), !counter) in + let (graph, node_l) = List.fold_right (fun sv (graph,node_l) -> + let (graph,node) = convert_sv_rec sv graph in + graph, node :: node_l + ) svl (graph,[]) in + let (graph, n) = convert_sv_rec sv graph in + let (graph, node_m) = convert_sm_rec sm graph in + let graph = S.add_vertex graph node in + let graph = List.fold_right (fun n graph -> + S.add_edge graph node n + ) node_l graph in + let graph = S.add_edge graph node n in + let graph = S.add_edge graph node node_m in + (graph,node) + +let convert_sv sv = convert_sv_rec sv S.empty +let convert_sm sm = convert_sm_rec sm S.empty + +let convert_sym st sm regs = + let graph = State.fold (fun res sv g -> + if (not (List.mem res regs)) then g + else + let (graph,head) = convert_sv sv in + incr counter; + let src = S.V.create (Oresource (Reg res), !counter) in + let graph = S.add_vertex graph src in + let graph = S.add_edge graph src head in + Op.union g graph + ) st S.empty + in + let graph' = + let (graph,head) = convert_sm sm in + incr counter; + let src = S.V.create (Oresource Mem, !counter) in + let graph = S.add_vertex graph src in + let graph = S.add_edge graph src head in + graph + in + Op.union graph graph' + +let display_st name l regs = + let (st,sm,_) = symbolic_evaluation initial_state initial_mem initial_cons l in + let g = convert_sym st sm regs in + let addr = SPDebug.name ^ name in + dot_output_ss g addr ; + ignore (Sys.command ("(dot -Tpng " ^ addr ^ " -o " ^ addr ^ ".png ; rm -f " ^ addr ^ ") & ")) (* & *) + + +let symbolic_equivalence (st1,sm1,cs1) (st2,sm2,cs2) regs = + Printf.printf "|cs1| = %i - |cs2| = %i \n" (Cons.cardinal cs1) (Cons.cardinal cs2); + (List.fold_right (fun res b -> + find res st1 = find res st2 && b + ) regs true + && sm1 = sm2 + && Cons.equal cs1 cs2) + +let symbolic_evaluation = symbolic_evaluation initial_state initial_mem initial_cons + +let symbolic_condition i l = + match i with + | Icond (cond,args) -> + let args = to_caml_list args in + let (st,sm,cs) = symbolic_evaluation l in + (cond,List.map (fun r -> find r st) args) + | _ -> failwith "Not a condition !\n" diff --git a/src/pipelining/SPTyping.ml b/src/pipelining/SPTyping.ml new file mode 100644 index 0000000..9b9c679 --- /dev/null +++ b/src/pipelining/SPTyping.ml @@ -0,0 +1,526 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +(*open Datatypes +open List +open Camlcoq +open Maps +open AST +open Op +open Registers +open RTL + +open Conventions +open Coqlib +open Errors +open Specif + +exception Type_error of string + +let env = ref (PTree.empty : typ PTree.t) + +let set_type r ty = + match PTree.get r !env with + | None -> env := PTree.set r ty !env + | Some ty' -> if ty <> ty' then + begin + Printf.fprintf SPDebug.dc "Failed to type register : %i " (P.to_int r); + raise (Type_error "type mismatch") + end + +let rec set_types rl tyl = + match rl, tyl with + | [], [] -> () + | r1 :: rs, ty1 :: tys -> set_type r1 ty1; set_types rs tys + | _, _ -> raise (Type_error "arity mismatch") + +(* First pass: process constraints of the form typeof(r) = ty *) + +let type_instr retty (pc, i) = + Printf.fprintf SPDebug.dc "typage de l'instruction : %i \n" (P.to_int pc); + match i with + | Inop(_) -> + () + | Iop(Omove, _, _, _) -> + () + | Iop(op, args, res, _) -> + let (targs, tres) = type_of_operation op in + set_types args targs; set_type res tres + | Iload(chunk, addr, args, dst, _) -> + set_types args (type_of_addressing addr); + set_type dst (type_of_chunk chunk) + | Istore(chunk, addr, args, src, _) -> + set_types args (type_of_addressing addr); + set_type src (type_of_chunk chunk) + | Icall(sg, ros, args, res, _) -> + begin try + begin match ros with + | Coq_inl r -> set_type r Tint + | Coq_inr _ -> () + end; + set_types args sg.sig_args; + set_type res (match sg.sig_res with Tret t -> t | _ -> Tint); + with Type_error msg -> + let name = + match ros with + | Coq_inl _ -> "" + | Coq_inr id -> extern_atom id in + raise(Type_error (Printf.sprintf "type mismatch in Icall(%s): %s" + name msg)) + end + | Itailcall(sg, ros, args) -> + begin try + begin match ros with + | Coq_inl r -> set_type r Tint + | Coq_inr _ -> () + end; + set_types args sg.sig_args; + if sg.sig_res <> retty then + raise (Type_error "mismatch on return type") + with Type_error msg -> + let name = + match ros with + | Coq_inl _ -> "" + | Coq_inr id -> extern_atom id in + raise(Type_error (Printf.sprintf "type mismatch in Itailcall(%s): %s" + name msg)) + end +(* | Ialloc(arg, res, _) -> + set_type arg Tint; set_type res Tint*) + | Icond(cond, args, _, _) -> + set_types args (type_of_condition cond) + | Ireturn(optres) -> + begin match optres, retty with + | None, Tvoid -> () + | Some r, Tret ty -> set_type r ty + | _, _ -> raise (Type_error "type mismatch in Ireturn") + end + +let type_pass1 retty instrs = + List.iter (type_instr retty) instrs + +(* Second pass: extract move constraints typeof(r1) = typeof(r2) + and solve them iteratively *) + +let rec extract_moves = function + | [] -> [] + | (pc, i) :: rem -> + match i with + | Iop(Omove, [r1], r2, _) -> + (r1, r2) :: extract_moves rem + | Iop(Omove, _, _, _) -> + raise (Type_error "wrong Omove") + | _ -> + extract_moves rem + +let changed = ref false + +let rec solve_moves = function + | [] -> [] + | (r1, r2) :: rem -> + match (PTree.get r1 !env, PTree.get r2 !env) with + | Some ty1, Some ty2 -> + if ty1 = ty2 + then (changed := true; solve_moves rem) + else raise (Type_error "type mismatch in Omove") + | Some ty1, None -> + env := PTree.set r2 ty1 !env; changed := true; solve_moves rem + | None, Some ty2 -> + env := PTree.set r1 ty2 !env; changed := true; solve_moves rem + | None, None -> + (r1, r2) :: solve_moves rem + +let rec iter_solve_moves mvs = + changed := false; + let mvs' = solve_moves mvs in + if !changed then iter_solve_moves mvs' + +let type_pass2 instrs = + iter_solve_moves (extract_moves instrs) + +let typeof e r = + match PTree.get r e with Some ty -> ty | None -> Tint + +let infer_type_environment f instrs = + try + env := PTree.empty; + set_types f.fn_params f.fn_sig.sig_args; + type_pass1 f.fn_sig.sig_res instrs; + type_pass2 instrs; + let e = !env in + env := PTree.empty; + Some(typeof e) + with Type_error msg -> + Printf.eprintf "Error during RTL type inference: %s\n" msg; + None + +(** val typ_eq : typ -> typ -> bool **) + +let typ_eq t1 t2 = + match t1 with + | Tint -> (match t2 with + | Tint -> true + | Tfloat -> false) + | Tfloat -> (match t2 with + | Tint -> false + | Tfloat -> true) + +(** val opt_typ_eq : typ option -> typ option -> bool **) + +let opt_typ_eq t1 t2 = + match t1 with + | Some x -> (match t2 with + | Some t -> typ_eq x t + | None -> false) + | None -> (match t2 with + | Some t -> false + | None -> true) + +(** val check_reg : regenv -> reg -> typ -> bool **) + +let check_reg env r ty = + match typ_eq (env r) ty with + | true -> true + | false -> false + +(** val check_regs : regenv -> reg list -> typ list -> bool **) + +let rec check_regs env rl tyl = + match rl with + | [] -> + (match tyl with + | [] -> true + | (t :: l) -> false) + | r1 :: rs -> + (match tyl with + | [] -> false + | (ty :: tys) -> + (match typ_eq (env r1) ty with + | true -> check_regs env rs tys + | false -> false)) + +(** val check_op : regenv -> operation -> reg list -> reg -> bool **) + +let check_op env op args res0 = + let (targs, tres) = type_of_operation op in + (match check_regs env args targs with + | true -> + (match typ_eq (env res0) tres with + | true -> true + | false -> false) + | false -> false) + +(** val check_successor : coq_function -> node -> bool **) + +let check_successor funct s = + match Maps.PTree.get s funct.fn_code with + | Some i -> true + | None -> false + +(** val check_instr : coq_function -> regenv -> instruction -> bool **) + +let check_instr funct env = function + | Inop s -> check_successor funct s + | Iop (op, args, res0, s) -> + (match op with + | Omove -> + (match args with + | Coq_nil -> false + | Coq_cons (arg, l) -> + (match l with + | Coq_nil -> + (match typ_eq (env arg) (env res0) with + | true -> check_successor funct s + | false -> false) + | Coq_cons (r, l0) -> false)) + | Ointconst i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ofloatconst f -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) +(* | Oaddrsymbol (i0, i1) -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false)*) +(* | Oaddrstack i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false)*) + | Ocast8signed -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ocast8unsigned -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ocast16signed -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ocast16unsigned -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) +(* | Oadd -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oaddimm i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false)*) + | Osub -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Osubimm i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Omul -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Omulimm i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Odiv -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Odivu -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oand -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oandimm i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oor -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oorimm i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oxor -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oxorimm i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Onand -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Onor -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Onxor -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oshl -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oshr -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oshrimm i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oshrximm i0 -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oshru -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Orolm (i0, i1) -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Onegf -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oabsf -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Oaddf -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Osubf -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Omulf -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Odivf -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Omuladdf -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Omulsubf -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Osingleoffloat -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ointoffloat -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ointuoffloat -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ofloatofint -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ofloatofintu -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false) + | Ocmp c -> + (match check_op env op args res0 with + | true -> check_successor funct s + | false -> false)) + | Iload (chunk, addr, args, dst, s) -> + (match check_regs env args (type_of_addressing addr) with + | true -> + (match typ_eq (env dst) (type_of_chunk chunk) with + | true -> check_successor funct s + | false -> false) + | false -> false) + | Istore (chunk, addr, args, src, s) -> + (match check_regs env args (type_of_addressing addr) with + | true -> + (match typ_eq (env src) (type_of_chunk chunk) with + | true -> check_successor funct s + | false -> false) + | false -> false) + | Icall (sig0, ros, args, res0, s) -> + (match match ros with + | Coq_inl r -> + (match typ_eq (env r) Tint with + | true -> check_regs env args sig0.sig_args + | false -> false) + | Coq_inr s0 -> check_regs env args sig0.sig_args with + | true -> + (match typ_eq (env res0) (proj_sig_res sig0) with + | true -> check_successor funct s + | false -> false) + | false -> false) + | Itailcall (sig0, ros, args) -> + (match match match ros with + | Coq_inl r -> + (match typ_eq (env r) Tint with + | true -> check_regs env args sig0.sig_args + | false -> false) + | Coq_inr s -> check_regs env args sig0.sig_args with + | true -> + proj_sumbool + (opt_typ_eq sig0.sig_res funct.fn_sig.sig_res) + | false -> false with + | true -> tailcall_is_possible sig0 + | false -> false) +(* | Ialloc (arg, res0, s) -> + (match typ_eq (env arg) Tint with + | true -> + (match typ_eq (env res0) Tint with + | true -> check_successor funct s + | false -> false) + | false -> false)*) + | Icond (cond, args, s1, s2) -> + (match match check_regs env args (type_of_condition cond) with + | true -> check_successor funct s1 + | false -> false with + | true -> check_successor funct s2 + | false -> false) + | Ireturn optres -> + (match optres with + | Some r -> + (match funct.fn_sig.sig_res with + | Some t -> + (match typ_eq (env r) t with + | true -> true + | false -> false) + | None -> false) + | None -> + (match funct.fn_sig.sig_res with + | Some t -> false + | None -> true)) + +(** val check_params_norepet : reg list -> bool **) + +let check_params_norepet params = + match list_norepet_dec Registers.Reg.eq params with + | true -> true + | false -> false + +(** val check_instrs : coq_function -> regenv -> (node, instruction) prod + list -> bool **) + +let rec check_instrs funct env = function + | Coq_nil -> true + | Coq_cons (p, rem) -> + let Coq_pair (pc, i) = p in + (match check_instr funct env i with + | true -> check_instrs funct env rem + | false -> false) + +(** val type_function : coq_function -> unit **) + +let type_function f = + let instrs = Maps.PTree.elements f.fn_code in + match infer_type_environment f instrs with + | Some env -> + (match match match match check_regs env f.fn_params + f.fn_sig.sig_args with + | true -> check_params_norepet f.fn_params + | false -> false with + | true -> check_instrs f env instrs + | false -> false with + | true -> check_successor f f.fn_entrypoint + | false -> false with + | true -> Printf.fprintf SPDebug.dc "The code is well typed\n" + | false -> failwith "Type checking failure\n") + | None -> failwith "Type inference failure\n" +*) diff --git a/src/pipelining/SPTyping.mli b/src/pipelining/SPTyping.mli new file mode 100644 index 0000000..dd27875 --- /dev/null +++ b/src/pipelining/SPTyping.mli @@ -0,0 +1,14 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + + +(*val type_function : RTL.coq_function -> unit*) diff --git a/src/pipelining/SoftwarePipelining.ml b/src/pipelining/SoftwarePipelining.ml new file mode 100644 index 0000000..0ba6d9d --- /dev/null +++ b/src/pipelining/SoftwarePipelining.ml @@ -0,0 +1,74 @@ +(***********************************************************************) +(* *) +(* Compcert Extensions *) +(* *) +(* Jean-Baptiste Tristan *) +(* *) +(* All rights reserved. This file is distributed under the terms *) +(* described in file ../../LICENSE. *) +(* *) +(***********************************************************************) + + +open SPBasic +open SPIMS +open SPMVE +open RTL + +let clean t = + + let rec clean_rec i = + match i with + | 0 -> [] + | n -> + begin + match t.(i - 1) with + | None -> clean_rec (i - 1) + | Some inst -> inst :: clean_rec (i - 1) + end + in + let l = List.rev (clean_rec (Array.length t)) in + List.hd l :: (List.filter (fun e -> not (is_cond e)) (List.tl l)) + +let print_nodes = List.iter (fun n -> Printf.printf "%s \n" (string_of_node n)) + +(* random heuristic *) + +let find node schedule opt = + try NI.find node schedule with + | Not_found -> opt + +(* A random heuristic is used to pick the next instruction to be scheduled from the unscheduled + * instructions. The scheduled instructions are given to the function, and the unscheduled + * instructions are created by taking all the instructions that are not in the scheduled list. + *) +let random ddg schedule = + let unscheduled = G.fold_vertex (fun node l -> + match find node schedule None with + | Some v -> l + | None -> node :: l + ) ddg [] in + let bound = List.length unscheduled in + Random.self_init (); + List.nth unscheduled (Random.int bound) + +(* tought heuristics *) + +module Topo = Graph.Topological.Make (G) +module Scc = Graph.Components.Make (G) + +let order = ref [] + +let pipeliner ddg = + order := List.flatten (Scc.scc_list ddg); + let (sched,ii) = SPIMS.pipeliner ddg random in + let (steady,prolog,epilog,min,unroll,entrance,way_out) = SPMVE.mve ddg sched ii in + let steady_state = clean steady in + if min <= 0 then None + else + Some {steady_state = steady_state; prolog = prolog; epilog = epilog; min = min; unrolling = unroll; + ramp_up = entrance; ramp_down = way_out} + + +let pipeline f = + SPBasic.apply_pipeliner f pipeliner ~debug:false -- cgit From e6348c97faee39754efd13b69a70c54851e2a789 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:11:02 +0000 Subject: Add bourdoncle code --- src/bourdoncle/Bourdoncle.v | 9 ++ src/bourdoncle/BourdoncleAux.ml | 109 +++++++++++++++++++++ src/bourdoncle/README.md | 3 + src/bourdoncle/bourdoncleIterator.ml | 183 +++++++++++++++++++++++++++++++++++ 4 files changed, 304 insertions(+) create mode 100644 src/bourdoncle/Bourdoncle.v create mode 100644 src/bourdoncle/BourdoncleAux.ml create mode 100644 src/bourdoncle/README.md create mode 100644 src/bourdoncle/bourdoncleIterator.ml diff --git a/src/bourdoncle/Bourdoncle.v b/src/bourdoncle/Bourdoncle.v new file mode 100644 index 0000000..9d470b7 --- /dev/null +++ b/src/bourdoncle/Bourdoncle.v @@ -0,0 +1,9 @@ +(** Type of a Bourdoncle component. *) + +Require Import List. +Require Import BinPos. +Notation node := positive. + +Inductive bourdoncle := +| I : node -> bourdoncle +| L : node -> list bourdoncle -> bourdoncle. diff --git a/src/bourdoncle/BourdoncleAux.ml b/src/bourdoncle/BourdoncleAux.ml new file mode 100644 index 0000000..3c7fd58 --- /dev/null +++ b/src/bourdoncle/BourdoncleAux.ml @@ -0,0 +1,109 @@ +(* Note: this file is used by SliceToString, so it must not include it *) + +open Camlcoq +open BinNums +open BinPos +open Maps +open RTL +open BourdoncleIterator +module B = Bourdoncle + +type node = P.t + +let int_of_positive p = + let i = P.to_int p in + i - 1 + +let positive_of_int n = P.of_int (n+1) + +(* Functions copied from SliceToString to avoid mutual inclusion *) +let nodeToString (p : P.t) : string = + Int.to_string (P.to_int p) +let rec cleanListToString' (aToString: 'a -> string) (l : 'a list) = + match l with + | [] -> "" + | e :: r -> " " ^ (aToString e) ^ (cleanListToString' aToString r) +let cleanListToString (aToString: 'a -> string) (l : 'a list) = + match l with + | [] -> "[]" + | [e] -> "(" ^ (aToString e) ^ ")" + | e :: r -> "(" ^ (aToString e) ^ (cleanListToString' aToString r) ^ ")" +let rec bourdoncleToString (b : B.bourdoncle) : string = + match b with + | B.I n -> (nodeToString n) + | B.L (h, lb) -> cleanListToString bourdoncleToString ((B.I h) :: lb) +let bourdoncleListToString (l : B.bourdoncle list) : string = + cleanListToString bourdoncleToString l + +(* Dummy type to avoid redefining existing functions *) +type instr = | Inop + +let build_cfg f = + let entry = int_of_positive f.fn_entrypoint in + let max = PTree.fold (fun m n _ -> max m (int_of_positive n)) f.fn_code 0 in + (* nodes are between 1 and max+1 *) + let succ = Array.make (max+1) [] in + let _ = PTree.fold (fun () n ins -> + succ.(int_of_positive n) <- + let inop : instr = Inop in + match ins with + | RTL.Inop j -> [(int_of_positive j, inop)] + | RTL.Iop (op,args,dst,j) -> [(int_of_positive j, Inop)] + | RTL.Iload (_,_,_,dst,j) -> [(int_of_positive j, Inop)] + | RTL.Istore (_,_,_,_,j) -> [(int_of_positive j, Inop)] + | RTL.Icall (_,_,_,dst,j) -> [(int_of_positive j, Inop)] + | RTL.Itailcall _ -> [] + | RTL.Ibuiltin (_,_,dst,j) -> [(int_of_positive j, Inop)] + | RTL.Icond (c,args,j,k) -> [(int_of_positive j, Inop);(int_of_positive k, Inop)] + | RTL.Ijumptable (_,tbl) -> List.map (fun j -> (int_of_positive j, inop)) tbl + | RTL.Ireturn _ -> []) f.fn_code () in + { entry = entry; succ = succ } + +let rec build_bourdoncle' (bl : bourdoncle list) : B.bourdoncle list = + match bl with + | [] -> [] + | (I i) :: r -> B.I (positive_of_int i) :: (build_bourdoncle' r) + | (L ((I h) :: l, _)) :: r -> B.L (positive_of_int h, build_bourdoncle' l) :: (build_bourdoncle' r) + | _ -> failwith "Assertion error: invalid bourdoncle ist" + +let build_bourdoncle'' (f : coq_function) : B.bourdoncle = + let cfg = build_cfg f in + match build_bourdoncle' (get_bourdoncle cfg) with + | [] -> failwith "assertion error: empty bourdoncle" + | l -> + begin + match l with + | B.I h :: r -> B.L (h, r) + | B.L (_, _) :: _ -> + begin + Printf.printf "ASSERTION ERROR: invalid program structure (too many bourdoncles)\n"; + Printf.printf "Head should be an element, but it is a list\n"; + Printf.printf "Failed at: %s\n" (bourdoncleListToString l); + failwith "assertion error" + end + | _ -> failwith "assertion error : ???" + end + +(* Auxiliary function for build_order *) +let rec linearize (b : B.bourdoncle) : node list = + match b with + | B.I n -> [n] + | B.L (h, l) -> List.fold_left (fun l0 b' -> l0 @ (linearize b')) [h] l + +let succ_pos = function + | N0 -> Npos Coq_xH + | Npos p -> Npos (Pos.succ p) + +let rec build_order' (l : node list) (count : coq_N) : coq_N PTree.t = + match l with + | [] -> PTree.empty + | n :: r -> PTree.set n count (build_order' r (succ_pos count)) + +let build_order (b : B.bourdoncle) : coq_N PMap.t = + let bo = build_order' (linearize b) (succ_pos N0) in + (succ_pos N0, bo) + +let build_bourdoncle (f : coq_function) : (B.bourdoncle * coq_N PMap.t) = + let b = build_bourdoncle'' f in + let bo = build_order b in + (b, bo) diff --git a/src/bourdoncle/README.md b/src/bourdoncle/README.md new file mode 100644 index 0000000..e1d02b9 --- /dev/null +++ b/src/bourdoncle/README.md @@ -0,0 +1,3 @@ +# Bourdoncle implementation + +This code was taken from WCET analysis in CompCert, which can be found [here](http://www.irisa.fr/celtique/ext/loopbound/). diff --git a/src/bourdoncle/bourdoncleIterator.ml b/src/bourdoncle/bourdoncleIterator.ml new file mode 100644 index 0000000..3cf0c78 --- /dev/null +++ b/src/bourdoncle/bourdoncleIterator.ml @@ -0,0 +1,183 @@ +type ('ab, 'instr) adom = { + order: 'ab -> 'ab -> bool; + join: 'ab -> 'ab -> 'ab; + widen: 'ab -> 'ab -> 'ab; + top: unit -> 'ab; + bot: unit -> 'ab; + transfer: 'instr -> 'ab -> 'ab; + to_string: 'ab -> string +} + +type 'instr cfg = { + entry: int; + succ: (int * 'instr) list array; +} + +(* Utilities *) +let rec print_list_sep_rec sep pp = function + | [] -> "" + | x::q -> sep^(pp x)^(print_list_sep_rec sep pp q) + +let rec print_list_sep_list_rec sep pp = function + | [] -> [] + | x::q -> (sep^(pp x))::(print_list_sep_list_rec sep pp q) + +let print_list_sep sep pp = function + | [] -> "" + | x::q -> (pp x)^(print_list_sep_rec sep pp q) + +let dot_cfg cfg filename = + let f = open_out filename in + Printf.fprintf f "digraph {\n"; + Array.iteri + (fun i l -> + List.iter (fun (j,_) -> Printf.fprintf f " n%d -> n%d\n" i j) l) + cfg.succ; + Printf.fprintf f "}\n"; + close_out f + + +(* Bourdoncle *) +type bourdoncle = L of (bourdoncle list) * Ptset.t | I of int + +let rec string_of_bourdoncle_list l = + print_list_sep " " string_of_bourdoncle l +and string_of_bourdoncle = function + | L (l,_) -> "("^(string_of_bourdoncle_list l)^")" + | I i -> string_of_int i + +(* For efficiency we pre-compute the set of nodes in a cfc *) +let add_component = + List.fold_left + (fun s b -> + match b with + | I i -> Ptset.add i s + | L (_,s') -> Ptset.union s s') + Ptset.empty + +(* Bourdoncle strategy computation *) +let get_bourdoncle cfg = + let date = ref 0 in + let debut = Array.map (fun _ -> 0) cfg.succ in + let stack = Stack.create () in + let rec component i = + let partition = ref [] in + List.iter + (fun (j,_) -> + if debut.(j)=0 then ignore (visite j partition)) + cfg.succ.(i); + (I i) :: !partition + and visite i partition = + incr date; + debut.(i) <- !date; + let min = ref !date in + let loop = ref false in + Stack.push i stack; + List.iter + (fun (j,_) -> + let m = if debut.(j)=0 then visite j partition else debut.(j) in + if m <= !min then + begin + min := m; + loop := true + end) + cfg.succ.(i); + if !min=debut.(i) then + begin + debut.(i) <- max_int; + let k = ref (Stack.pop stack) in + if !loop then begin + while (!k<>i) do + debut.(!k) <- 0; + k := Stack.pop stack; + done; + let c = component i in + partition := (L (c,add_component c)) :: !partition + end + else partition := (I i) :: !partition + end; + !min + in + let partition = ref [] in + (* only one entry in the cfg *) + ignore (visite cfg.entry partition); + Array.iteri (fun i d -> if d=0 then cfg.succ.(i) <- []) debut; + !partition + +(* predecessors *) +let get_pred cfg = + let pred = Array.make (Array.length cfg.succ) [] in + Array.iteri + (fun i succs -> + List.iter + (fun (j,ins) -> + pred.(j) <- (i,ins) :: pred.(j)) + succs + ) + cfg.succ; + pred + +let check_fixpoint adom cfg res = + if not (adom.order (adom.top ()) res.(cfg.entry)) then + failwith "res.(ENTRY) should be top !\n"; + Array.iteri + (fun i succs -> + List.iter + (fun (j,op) -> + if not (adom.order (adom.transfer op res.(i)) res.(j)) then + begin + dot_cfg cfg "debug_bourdoncle.dot"; + failwith (Printf.sprintf "res.(%d) should be lower than res.(%d) !\n" i j) + end) + succs) + cfg.succ + +let option_down_iterations = ref 1 + +let rec f_list f acc = function + [] -> acc + | x::q -> f_list f (f acc x) q + +let run_with_bourdoncle adom b cfg = + let preds = get_pred cfg in + let res = Array.init (Array.length preds) (fun _ -> adom.bot ()) in + let join_list = f_list adom.join in + let _ = res.(cfg.entry) <- adom.top () in + let upd j = + if j <> cfg.entry then + res.(j) <- + join_list (adom.bot ()) + (List.map (fun (i,op) -> adom.transfer op res.(i)) preds.(j)) in + let upd_except_cfc j cfc = + if j <> cfg.entry then + res.(j) <- join_list res.(j) + (List.map (fun (i,op) -> adom.transfer op res.(i)) + (List.filter (fun (i,op) -> not (Ptset.mem i cfc)) preds.(j))) + in + let nb_down = !option_down_iterations in + let rec iter_component down = function + | L (I head::rest,cfc) -> + let rec aux down = function + [] -> + let old_ab = res.(head) in + let _ = upd head in + let new_ab = res.(head) in + if down >= nb_down || (new_ab = old_ab) then () + else if not (adom.order new_ab old_ab) then + begin + res.(head) <- adom.widen old_ab new_ab; + aux (-1) rest + end + else (upd head; aux (down+1) rest) + | b::q -> iter_component down b; aux down q in + upd_except_cfc head cfc; + aux down rest + | I i -> upd i + | _ -> assert false in + List.iter (iter_component (-1)) b; + check_fixpoint adom cfg res; + res + +let run adom cfg = + let b = get_bourdoncle cfg in + run_with_bourdoncle adom b cfg -- 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 --- scripts/synth.tcl | 2 +- src/Compiler.v | 20 ++++---- src/extraction/Extraction.v | 2 +- 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 ++--- 12 files changed, 163 insertions(+), 161 deletions(-) diff --git a/scripts/synth.tcl b/scripts/synth.tcl index e5151e8..a2fb722 100644 --- a/scripts/synth.tcl +++ b/scripts/synth.tcl @@ -76,7 +76,7 @@ proc dump_statistics { } { }; #END PROC set outputDir . create_project -in_memory -part xc7z020clg484-1 -force -read_verilog main.v +read_verilog -sv main.v synth_design -mode out_of_context -no_iobuf -top main -part xc7z020clg484-1 write_checkpoint -force $outputDir/post_synth.dcp report_timing_summary -file $outputDir/post_synth_timing_summary.rpt diff --git a/src/Compiler.v b/src/Compiler.v index 0ac2b80..e4ab607 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -196,8 +196,8 @@ Definition transf_backend (r : RTL.program) : res Verilog.program := @@ print (print_RTL 7) @@@ HTLgen.transl_program @@ print (print_HTL 0) - @@ total_if HLSOpts.optim_ram Memorygen.transf_program - @@ print (print_HTL 1) +(* @@ total_if HLSOpts.optim_ram Memorygen.transf_program + @@ print (print_HTL 1)*) @@ Veriloggen.transl_program. (*| @@ -279,7 +279,7 @@ Definition CompCert's_passes := ::: mkpass (match_if Compopts.optim_redundancy Deadcodeproof.match_prog) ::: mkpass Unusedglobproof.match_prog ::: (@mkpass _ _ HTLgenproof.match_prog (HTLgenproof.TransfHTLLink HTLgen.transl_program)) - ::: mkpass (match_if HLSOpts.optim_ram Memorygen.match_prog) + (*::: mkpass (match_if HLSOpts.optim_ram Memorygen.match_prog)*) ::: mkpass Veriloggenproof.match_prog ::: pass_nil _. @@ -317,8 +317,8 @@ Proof. destruct (partial_if Compopts.optim_redundancy Deadcode.transf_program p11) as [p12|e] eqn:P12; simpl in T; try discriminate. destruct (Unusedglob.transform_program p12) as [p13|e] eqn:P13; simpl in T; try discriminate. destruct (HTLgen.transl_program p13) as [p14|e] eqn:P14; simpl in T; try discriminate. - set (p15 := total_if HLSOpts.optim_ram Memorygen.transf_program p14) in *. - set (p16 := Veriloggen.transl_program p15) in *. + (*set (p15 := total_if HLSOpts.optim_ram Memorygen.transf_program p14) in *.*) + set (p15 := Veriloggen.transl_program p14) in *. unfold match_prog; simpl. exists p1; split. apply SimplExprproof.transf_program_match; auto. exists p2; split. apply SimplLocalsproof.match_transf_program; auto. @@ -334,8 +334,8 @@ Proof. exists p12; split. eapply partial_if_match; eauto. apply Deadcodeproof.transf_program_match. exists p13; split. apply Unusedglobproof.transf_program_match; auto. exists p14; split. apply HTLgenproof.transf_program_match; auto. - exists p15; split. apply total_if_match. apply Memorygen.transf_program_match; auto. - exists p16; split. apply Veriloggenproof.transf_program_match; auto. + (*exists p15; split. apply total_if_match. apply Memorygen.transf_program_match; auto.*) + exists p15; split. apply Veriloggenproof.transf_program_match; auto. inv T. reflexivity. Qed. @@ -353,7 +353,7 @@ Ltac DestructM := destruct H as (p & M & MM); clear H end. repeat DestructM. subst tp. - assert (F: forward_simulation (Cstrategy.semantics p) (Verilog.semantics p16)). + assert (F: forward_simulation (Cstrategy.semantics p) (Verilog.semantics p15)). { eapply compose_forward_simulations. eapply SimplExprproof.transl_program_correct; eassumption. @@ -383,8 +383,8 @@ Ltac DestructM := eapply Unusedglobproof.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply HTLgenproof.transf_program_correct. eassumption. - eapply compose_forward_simulations. - eapply match_if_simulation. eassumption. exact Memorygen.transf_program_correct; eassumption. + (*eapply compose_forward_simulations. + eapply match_if_simulation. eassumption. exact Memorygen.transf_program_correct; eassumption.*) eapply Veriloggenproof.transf_program_correct; eassumption. } split. auto. diff --git a/src/extraction/Extraction.v b/src/extraction/Extraction.v index aeefd2a..9b82e1b 100644 --- a/src/extraction/Extraction.v +++ b/src/extraction/Extraction.v @@ -182,7 +182,7 @@ Extract Inlined Constant Binary.B2R => "fun _ -> assert false". Extract Inlined Constant Binary.round_mode => "fun _ -> assert false". Extract Inlined Constant Bracket.inbetween_loc => "fun _ -> assert false". -Extract Constant Pipeline.pipeline => "SoftwarePipelining.pipeline". +Extract Constant Pipeline.pipeline => "pipelining.pipeline". Extract Constant RTLBlockgen.partition => "Partition.partition". Extract Constant RTLPargen.schedule => "Schedule.schedule_fn". 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 90ed9c794dadc04259530f7058cd9bc260814d33 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:14:43 +0000 Subject: Improve the benchmark Makefile --- benchmarks/polybench-syn/common.mk | 2 +- benchmarks/polybench-syn/run-vericert.sh | 10 +++++++--- 2 files changed, 8 insertions(+), 4 deletions(-) diff --git a/benchmarks/polybench-syn/common.mk b/benchmarks/polybench-syn/common.mk index 88fb059..85d62f6 100644 --- a/benchmarks/polybench-syn/common.mk +++ b/benchmarks/polybench-syn/common.mk @@ -1,5 +1,5 @@ VERICERT ?= vericert -VERICERT_OPTS ?= -DSYNTHESIS -O0 -finline +VERICERT_OPTS ?= -DSYNTHESIS -O0 -finline -fschedule -fif-conv IVERILOG ?= iverilog IVERILOG_OPTS ?= diff --git a/benchmarks/polybench-syn/run-vericert.sh b/benchmarks/polybench-syn/run-vericert.sh index ef6964f..9deaa10 100755 --- a/benchmarks/polybench-syn/run-vericert.sh +++ b/benchmarks/polybench-syn/run-vericert.sh @@ -10,27 +10,31 @@ while read benchmark ; do cresult=$(cat $benchmark.clog | cut -d' ' -f2) #echo "C output: "$cresult #./$benchmark.iver > $benchmark.tmp + if [[ ! -f ./$benchmark.verilator/Vmain ]]; then + echo -e "\e[0;91mFAIL\e[0m: Verilog failed compilation" + continue + fi ./$benchmark.verilator/Vmain > $benchmark.tmp veriresult=$(tail -1 $benchmark.tmp | cut -d' ' -f2) cycles=$(tail -2 $benchmark.tmp | head -1 | tr -s ' ' | cut -d' ' -f2) #echo "Verilog output: "$veriresult #Undefined checks - if test -z $veriresult + if [[ -z "$veriresult" ]] then echo "\e[0;91mFAIL\e[0m: Verilog returned nothing" #exit 0 fi # Don't care checks - if [ $veriresult == "x" ] + if [[ $veriresult == "x" ]] then echo "\e[0;91mFAIL\e[0m: Verilog returned don't cares" #exit 0 fi # unequal result check - if [ $cresult -ne $veriresult ] + if [[ $cresult -ne $veriresult ]] then echo -e "\e[0;91mFAIL\e[0m: Verilog and C output do not match!" #exit 0 -- cgit From fc215b1307b9b8a8c7392b7faf252fede26c7a35 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:15:02 +0000 Subject: Fix the pipelining link in the README --- README.org | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/README.org b/README.org index cf8bd34..6f50560 100644 --- a/README.org +++ b/README.org @@ -137,7 +137,8 @@ This project is licensed under [[https://www.gnu.org/licenses/gpl-3.0.en.html][G The following external code and its license is present in this repository: -- [[/src/SoftwarePipelining][/src/SoftwarePipelining]] :: MIT +- [[/src/pipelining +][/src/pipelining]] :: MIT #+begin_src text Copyright (c) 2008,2009,2010 Jean-Baptiste Tristan and INRIA -- 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(-) 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(-) 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 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 0deea69405d5b4380e874b5b0209bf8a54027972 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:16:28 +0000 Subject: Add a script to gather synthesis data --- scripts/gather_data.rkt | 66 +++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 66 insertions(+) create mode 100644 scripts/gather_data.rkt diff --git a/scripts/gather_data.rkt b/scripts/gather_data.rkt new file mode 100644 index 0000000..b2aa9c5 --- /dev/null +++ b/scripts/gather_data.rkt @@ -0,0 +1,66 @@ +#lang racket + +(require xml) +(require xml/path) +(require racket/match) + +(require csv-reading) +(require csv-writing) + +(require graphite) +(require data-frame) +(require sawzall) + +(require threading) + +(permissive-xexprs #t) + +(define (parse-vivado-report f) + (let* ([encode-xml-port (open-input-file f)] + [report (xml->xexpr (document-element (read-xml encode-xml-port)))]) + (close-input-port encode-xml-port) + report)) + +(define (list->hash l) + (foldr (lambda (v l) + (hash-set l (car v) (string->number (cadr v)))) + (hash) l)) + +(define (process-vivado-report report) + (let ([maps (map (lambda (x) (match x [(list e (list (list a b) (list c d))) (list b d)])) + (filter-not string? (se-path*/list '(section) report)))]) + (list->hash maps))) + +(define (parse-sim-report f) + (let* ([exec-csv (open-input-file f)] + [report (csv->list exec-csv)]) + (close-input-port exec-csv) + report)) + +(define sim-report (list->hash (parse-sim-report "exec.csv"))) +(define vivado-report (process-vivado-report (parse-vivado-report "encode_report.xml"))) +(define cycles (hash-ref sim-report "covariance")) +(define delay (hash-ref vivado-report "XILINX_DESIGN_DELAY")) + +(* cycles delay) + +(define gss (df-read/csv "exec.csv")) + +(df-add-series gss (make-series "tool" #:data (make-vector (df-row-count gss) "vericert"))) + +(df-add-series gss (make-series "delay" #:data (list->vector (map (lambda (x) + (~> "/encode_report.xml" + (string-append x _) + parse-vivado-report + process-vivado-report + (hash-ref "XILINX_DESIGN_DELAY"))) + (filter-not (lambda (x) (equal? x "test-case")) (map car (hash->list sim-report))))))) + +(show gss everything #:n-rows 'all) + +(graph #:data gss + #:title "Chart" + #:mapping (aes #:x "test-case" #:y "cycle count") + (col #:gap 0.25)) + +;(define csv (open-output-file "out.csv")) -- cgit From bbfba26e5131665dbf3e7f6adf0c1746dccec243 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 9 Dec 2021 15:36:03 +0000 Subject: Add README.md --- README.md | 147 ++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 147 insertions(+) create mode 100644 README.md diff --git a/README.md b/README.md new file mode 100644 index 0000000..baebe7a --- /dev/null +++ b/README.md @@ -0,0 +1,147 @@ +- [Features](#features) + - [Building](#building) + - [Downloading Vericert and CompCert](#downloading-compcert) + - [Setting up Nix](#setting-up-nix) + - [Makefile build](#makefile-build) + - [Running](#running) + - [Citation](#org741ad34) + - [License](#orgb089937) + + + +

 

+ +A formally verified high-level synthesis (HLS) tool written in Coq, building on top of [CompCert](https://github.com/AbsInt/CompCert). This ensures the correctness of the C to Verilog translation according to our Verilog semantics and CompCert’s C semantics, removing the need to check the resulting hardware for behavioural correctness. + + + + +# Features + +Currently all proofs of the following features have been completed. + +- all int operations, +- non-recursive function calls, +- local arrays and pointers +- control-flow structures such as if-statements, for-loops, etc… + + + + +# Building + +To build Vericert, the provided [Makefile](file:///Makefile) can be used. External dependencies are needed to build the project, which can be pulled in automatically with [nix](https://nixos.org/nix/) using the provided [default.nix](file:///default.nix) and [shell.nix](file:///shell.nix) files. + +The project is written in Coq, a theorem prover, which is extracted to OCaml so that it can then be compiled and executed. The dependencies of this project are the following: + +- [Coq](https://coq.inria.fr/): theorem prover that is used to also program the HLS tool. +- [OCaml](https://ocaml.org/): the OCaml compiler to compile the extracted files. +- [dune](https://github.com/ocaml/dune): build tool for ocaml projects to gather all the ocaml files and compile them in the right order. +- [menhir](http://gallium.inria.fr/~fpottier/menhir/): parser generator for ocaml. +- [findlib](https://github.com/ocaml/ocamlfind) to find installed OCaml libraries. +- [GCC](https://gcc.gnu.org/): compiler to help build CompCert. + +These dependencies can be installed manually, or automatically through Nix. + + + + +## Downloading Vericert and CompCert + +CompCert is added as a submodule in the `lib/CompCert` directory. It is needed to run the build process below, as it is the one dependency that is not downloaded by nix, and has to be downloaded together with the repository. To clone CompCert together with this project, and check it out at the correct revision, you can run: + +```shell +git clone -b v1.2.2 --recursive https://github.com/ymherklotz/vericert +``` + +If the repository is already cloned, you can run the following command to make sure that CompCert is also downloaded and the correct branch is checked out: + +```shell +git checkout v1.2.2 +git submodule update --init +``` + + + + +## Setting up Nix + +Nix is a package manager that can create an isolated environment so that the builds are reproducible. Once nix is installed, it can be used in the following way. + +To open a shell which includes all the necessary dependencies, one can use: + +```shell +nix-shell +``` + +which will open a shell that has all the dependencies loaded. + + + + +## Makefile build + +If the dependencies were installed manually, or if one is in the `nix-shell`, the project can be built by running: + +```shell +make -j8 +``` + +and installed locally, or under the `PREFIX` location using: + +```shell + make install +``` + +Which will install the binary in `./bin/vericert` by default. However, this can be changed by changing the `PREFIX` environment variable, in which case the binary will be installed in `$PREFIX/bin/vericert`. + + + + +# Running + +To test out `vericert` you can try the following examples which are in the test folder using the following: + +```shell +./bin/vericert test/loop.c -o loop.v +./bin/vericert test/conditional.c -o conditional.v +./bin/vericert test/add.c -o add.v +``` + + + + +# Citation + +If you use Vericert in any way, please cite it using our [OOPSLA’21 paper](https://yannherklotz.com/papers/fvhls_oopsla21.pdf): + +```bibtex +@inproceedings{herklotz21_fvhls, + author = {Herklotz, Yann and Pollard, James D. and Ramanathan, Nadesh and Wickerson, John}, + title = {Formal Verification of High-Level Synthesis}, + year = {2021}, + number = {OOPSLA}, + numpages = {30}, + month = {11}, + journal = {Proc. ACM Program. Lang.}, + volume = {5}, + publisher = {Association for Computing Machinery}, + address = {New York, NY, USA}, + doi = {10.1145/3485494} +} +``` + + + + +# License + +This project is licensed under [GPLv3](https://www.gnu.org/licenses/gpl-3.0.en.html). The license can be seen in [/LICENSE](file:///LICENSE). + +The following external code and its license is present in this repository: + +- **[/src/pipelining](file:///src/pipelining):** MIT + +```text +Copyright (c) 2008,2009,2010 Jean-Baptiste Tristan and INRIA +``` -- cgit From 048e219468508ad403e576ef15a416fd8d051b22 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 9 Dec 2021 19:56:42 +0000 Subject: Update README with correct link --- README.md | 8 ++++---- README.org | 3 +-- 2 files changed, 5 insertions(+), 6 deletions(-) diff --git a/README.md b/README.md index baebe7a..e801149 100644 --- a/README.md +++ b/README.md @@ -4,8 +4,8 @@ - [Setting up Nix](#setting-up-nix) - [Makefile build](#makefile-build) - [Running](#running) - - [Citation](#org741ad34) - - [License](#orgb089937) + - [Citation](#orge392203) + - [License](#org8a5a598) @@ -109,7 +109,7 @@ To test out `vericert` you can try the following examples which are in the test ``` - + # Citation @@ -132,7 +132,7 @@ If you use Vericert in any way, please cite it using our [OOPSLA’21 paper] ``` - + # License diff --git a/README.org b/README.org index 6f50560..da29bb1 100644 --- a/README.org +++ b/README.org @@ -137,8 +137,7 @@ This project is licensed under [[https://www.gnu.org/licenses/gpl-3.0.en.html][G The following external code and its license is present in this repository: -- [[/src/pipelining -][/src/pipelining]] :: MIT +- [[/src/pipelining][/src/pipelining]] :: MIT #+begin_src text Copyright (c) 2008,2009,2010 Jean-Baptiste Tristan and INRIA -- cgit From 1c91db6b6d4cd20994fdfc40283b003c1ec93d50 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 9 Dec 2021 20:02:36 +0000 Subject: Update benchmark runs --- benchmarks/polybench-syn/common.mk | 2 +- scripts/gather_data.rkt | 112 ++++++++++++++++++++++++++----------- 2 files changed, 80 insertions(+), 34 deletions(-) diff --git a/benchmarks/polybench-syn/common.mk b/benchmarks/polybench-syn/common.mk index 85d62f6..8e67294 100644 --- a/benchmarks/polybench-syn/common.mk +++ b/benchmarks/polybench-syn/common.mk @@ -1,5 +1,5 @@ VERICERT ?= vericert -VERICERT_OPTS ?= -DSYNTHESIS -O0 -finline -fschedule -fif-conv +VERICERT_OPTS ?= -DSYNTHESIS -fschedule -fif-conv IVERILOG ?= iverilog IVERILOG_OPTS ?= diff --git a/scripts/gather_data.rkt b/scripts/gather_data.rkt index b2aa9c5..5e97c9e 100644 --- a/scripts/gather_data.rkt +++ b/scripts/gather_data.rkt @@ -15,21 +15,62 @@ (permissive-xexprs #t) -(define (parse-vivado-report f) - (let* ([encode-xml-port (open-input-file f)] - [report (xml->xexpr (document-element (read-xml encode-xml-port)))]) - (close-input-port encode-xml-port) - report)) +(define (to-name f) (regexp-replace #rx".*/([^/]+)/encode_report.xml" f "\\1")) + +(define files + '("./data/data-mining/covariance/encode_report.xml" + "./data/stencils/heat-3d/encode_report.xml" + "./data/stencils/jacobi-1d/encode_report.xml" + "./data/stencils/seidel-2d/encode_report.xml" + "./data/stencils/jacobi-2d/encode_report.xml" + "./data/linear-algebra/kernels/doitgen/encode_report.xml" + "./data/linear-algebra/kernels/2mm/encode_report.xml" + "./data/linear-algebra/kernels/3mm/encode_report.xml" + "./data/linear-algebra/blas/gemver/encode_report.xml" + "./data/linear-algebra/blas/syrk/encode_report.xml" + "./data/linear-algebra/blas/gemm/encode_report.xml" + "./data/linear-algebra/solvers/trisolv/encode_report.xml")) (define (list->hash l) + (foldr (lambda (v l) + (hash-set l (car v) (cadr v))) + (hash) l)) + +(define (list->hashn l) (foldr (lambda (v l) (hash-set l (car v) (string->number (cadr v)))) (hash) l)) +(define name-f-map + (list->hash (map (lambda (f) (list (to-name f) f)) files))) + +(define (parse-vivado-report f) + (let* ([encode-xml-port (open-input-file f)] + [report (xml->xexpr (document-element (read-xml encode-xml-port)))]) + (close-input-port encode-xml-port) + report)) + (define (process-vivado-report report) (let ([maps (map (lambda (x) (match x [(list e (list (list a b) (list c d))) (list b d)])) (filter-not string? (se-path*/list '(section) report)))]) - (list->hash maps))) + (list->hashn maps))) + +(define (vivado-report-f f) (process-vivado-report (parse-vivado-report f))) + +;;(vivado-report-f "./data/data-mining/covariance/encode_report.xml") + +(define synth-f (map flatten (hash-map name-f-map + (lambda (n f) (let ([x (vivado-report-f f)]) + (list n + (hash-ref x "XILINX_LUT_FLIP_FLOP_PAIRS_USED") + (hash-ref x "XILINX_SLICE") + (hash-ref x "XILINX_SLICE_REGISTERS") + (hash-ref x "XILINX_SLICE_LUTS") + (hash-ref x "XILINX_BLOCK_RAMFIFO") + (hash-ref x "XILINX_IOPIN") + (hash-ref x "XILINX_DSPS") + (hash-ref x "XILINX_POWER") + (hash-ref x "XILINX_DESIGN_DELAY"))))))) (define (parse-sim-report f) (let* ([exec-csv (open-input-file f)] @@ -37,30 +78,35 @@ (close-input-port exec-csv) report)) -(define sim-report (list->hash (parse-sim-report "exec.csv"))) -(define vivado-report (process-vivado-report (parse-vivado-report "encode_report.xml"))) -(define cycles (hash-ref sim-report "covariance")) -(define delay (hash-ref vivado-report "XILINX_DESIGN_DELAY")) - -(* cycles delay) - -(define gss (df-read/csv "exec.csv")) - -(df-add-series gss (make-series "tool" #:data (make-vector (df-row-count gss) "vericert"))) - -(df-add-series gss (make-series "delay" #:data (list->vector (map (lambda (x) - (~> "/encode_report.xml" - (string-append x _) - parse-vivado-report - process-vivado-report - (hash-ref "XILINX_DESIGN_DELAY"))) - (filter-not (lambda (x) (equal? x "test-case")) (map car (hash->list sim-report))))))) - -(show gss everything #:n-rows 'all) - -(graph #:data gss - #:title "Chart" - #:mapping (aes #:x "test-case" #:y "cycle count") - (col #:gap 0.25)) - -;(define csv (open-output-file "out.csv")) +(define sim-report (list->hashn (parse-sim-report "exec.csv"))) + +(define csv (open-output-file "out2.csv")) +(display (table->string + (append '((benchmark lut_flip_flop slice regs luts ramfifo iopin dsps power delay cycles)) + (map (lambda (x) (append x (list (hash-ref sim-report (car x))))) synth-f))) csv) +(close-output-port csv) + +;;(define vivado-report (process-vivado-report (parse-vivado-report "encode_report.xml"))) +;;(define cycles (hash-ref sim-report "covariance")) +;;(define delay (hash-ref vivado-report "XILINX_DESIGN_DELAY")) +;; +;;(define gss (df-read/csv "exec.csv")) +;; +;;(df-add-series gss (make-series "tool" #:data (make-vector (df-row-count gss) "vericert"))) +;; +;;(df-add-series gss (make-series "delay" #:data (list->vector (map (lambda (x) +;; (~> "/encode_report.xml" +;; (string-append x _) +;; parse-vivado-report +;; process-vivado-report +;; (hash-ref "XILINX_DESIGN_DELAY"))) +;; (filter-not (lambda (x) (equal? x "test-case")) (map car (hash->list sim-report))))))) +;; +;;(show gss everything #:n-rows 'all) +;; +;;(graph #:data gss +;; #:title "Chart" +;; #:mapping (aes #:x "test-case" #:y "cycle count") +;; (col #:gap 0.25)) +;; +;;;(define csv (open-output-file "out.csv")) -- cgit From a64ccb64e16175b520a0dc4badde44a1cc46f17d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 9 Dec 2021 20:03:41 +0000 Subject: Remove debug directory --- debug/CoqupTest.ml | 59 ------------------------------------------------------ debug/dune | 5 ----- 2 files changed, 64 deletions(-) delete mode 100644 debug/CoqupTest.ml delete mode 100644 debug/dune diff --git a/debug/CoqupTest.ml b/debug/CoqupTest.ml deleted file mode 100644 index 2fe1389..0000000 --- a/debug/CoqupTest.ml +++ /dev/null @@ -1,59 +0,0 @@ -open Vericert - -open Test -open Camlcoq -open FMapPositive - -let test_function_transf () = - print_endline "Testing transformation"; - print_endline "TL PROGRAM: "; - PrintRTL.print_program stdout testinputprogram; - print_endline "VERILOG PROGRAM: "; - - let v = match Veriloggen.transf_program testinputprogram with - | Errors.OK v -> v - | Errors.Error _ -> - raise (Failure "Error") in - PrintVerilog.print_program stdout v - -let cvalue n = Value.natToValue (Nat.of_int 32) (Nat.of_int n) - -let test_values () = - print_endline "Testing values"; - let v1 = cvalue 138 in - let v2 = cvalue 12 in - let v3 = Value.natToValue (Nat.of_int 16) (Nat.of_int 21) in - PrintVerilog.print_value stdout v1; - print_newline(); - PrintVerilog.print_value stdout v2; - print_newline(); - PrintVerilog.print_value stdout v3; - print_newline(); - print_string "Addition: "; - PrintVerilog.print_value stdout (Value.vplus v1 v2); - print_newline(); - print_string "Wrong addition: "; - PrintVerilog.print_value stdout (Value.vplus v1 v3); - print_newline(); - print_string "Opt addition: "; - match Value.vplus_opt v1 v2 with - | Some x -> begin - PrintVerilog.print_value stdout x; - print_endline ""; - match Value.vplus_opt v1 v3 with - | Some x -> PrintVerilog.print_value stdout x; print_newline() - | None -> print_endline "Correctly failed in addition" - end - | None -> raise (Failure "Error") - -let simulate_test () = - print_endline "Simulation test"; - let v = - match Veriloggen.transf_program testinputprogram with - | Errors.OK v -> v - | _ -> raise (Failure "HLS Error") in - match Verilog.module_run (Nat.of_int 10) v with - | Errors.OK lst -> PrintVerilog.print_result stdout (PositiveMap.elements lst) - | _ -> raise (Failure "Simulation error") - -let () = test_function_transf () diff --git a/debug/dune b/debug/dune deleted file mode 100644 index feee97c..0000000 --- a/debug/dune +++ /dev/null @@ -1,5 +0,0 @@ -(include_subdirs no) - -(executable - (name VericertTest) - (libraries vericert)) -- 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 --- Makefile | 2 +- src/Compiler.v | 2 +- src/bourdoncle/BourdoncleAux.ml | 31 +++++++++++++++---------------- src/extraction/Extraction.v | 10 +++++++--- src/hls/IfConversion.v | 36 ++++++++++++++++++++++++++++++++---- src/hls/Pipeline.v | 3 ++- src/hls/PipelineOp.v | 3 ++- src/hls/Schedule.ml | 5 ++++- 8 files changed, 64 insertions(+), 28 deletions(-) diff --git a/Makefile b/Makefile index 0749d1c..5577ba9 100644 --- a/Makefile +++ b/Makefile @@ -21,7 +21,7 @@ COQMAKE := $(COQBIN)coq_makefile COQDOCFLAGS := --no-lib-name -l -VS := src/Compiler.v src/Simulator.v src/HLSOpts.v $(foreach d, common hls, src/$(d)/*.v) +VS := src/Compiler.v src/Simulator.v src/HLSOpts.v $(foreach d, common hls bourdoncle, src/$(d)/*.v) PREFIX ?= . diff --git a/src/Compiler.v b/src/Compiler.v index e4ab607..b907629 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -67,7 +67,7 @@ Require vericert.hls.RTLBlockgen. Require vericert.hls.RTLPargen. Require vericert.hls.RTLParFUgen. Require vericert.hls.HTLPargen. -Require vericert.hls.Pipeline. +(*Require vericert.hls.Pipeline.*) Require vericert.hls.IfConversion. (*Require vericert.hls.PipelineOp.*) Require vericert.HLSOpts. diff --git a/src/bourdoncle/BourdoncleAux.ml b/src/bourdoncle/BourdoncleAux.ml index 3c7fd58..a35004e 100644 --- a/src/bourdoncle/BourdoncleAux.ml +++ b/src/bourdoncle/BourdoncleAux.ml @@ -4,7 +4,8 @@ open Camlcoq open BinNums open BinPos open Maps -open RTL +open RTLBlock +open RTLBlockInstr open BourdoncleIterator module B = Bourdoncle @@ -38,25 +39,23 @@ let bourdoncleListToString (l : B.bourdoncle list) : string = (* Dummy type to avoid redefining existing functions *) type instr = | Inop +let rec get_exits = function + | RBgoto j -> [(int_of_positive j, Inop)] + | RBcall (_,_,_,dst,j) -> [(int_of_positive j, Inop)] + | RBtailcall _ -> [] + | RBbuiltin (_,_,dst,j) -> [(int_of_positive j, Inop)] + | RBcond (c,args,j,k) -> [(int_of_positive j, Inop);(int_of_positive k, Inop)] + | RBjumptable (_,tbl) -> List.map (fun j -> (int_of_positive j, Inop)) tbl + | RBpred_cf (p, cf1, cf2) -> List.append (get_exits cf1) (get_exits cf2) + | RBreturn _ -> [] + let build_cfg f = let entry = int_of_positive f.fn_entrypoint in let max = PTree.fold (fun m n _ -> max m (int_of_positive n)) f.fn_code 0 in (* nodes are between 1 and max+1 *) let succ = Array.make (max+1) [] in let _ = PTree.fold (fun () n ins -> - succ.(int_of_positive n) <- - let inop : instr = Inop in - match ins with - | RTL.Inop j -> [(int_of_positive j, inop)] - | RTL.Iop (op,args,dst,j) -> [(int_of_positive j, Inop)] - | RTL.Iload (_,_,_,dst,j) -> [(int_of_positive j, Inop)] - | RTL.Istore (_,_,_,_,j) -> [(int_of_positive j, Inop)] - | RTL.Icall (_,_,_,dst,j) -> [(int_of_positive j, Inop)] - | RTL.Itailcall _ -> [] - | RTL.Ibuiltin (_,_,dst,j) -> [(int_of_positive j, Inop)] - | RTL.Icond (c,args,j,k) -> [(int_of_positive j, Inop);(int_of_positive k, Inop)] - | RTL.Ijumptable (_,tbl) -> List.map (fun j -> (int_of_positive j, inop)) tbl - | RTL.Ireturn _ -> []) f.fn_code () in + succ.(int_of_positive n) <- get_exits ins.bb_exit) f.fn_code () in { entry = entry; succ = succ } let rec build_bourdoncle' (bl : bourdoncle list) : B.bourdoncle list = @@ -66,7 +65,7 @@ let rec build_bourdoncle' (bl : bourdoncle list) : B.bourdoncle list = | (L ((I h) :: l, _)) :: r -> B.L (positive_of_int h, build_bourdoncle' l) :: (build_bourdoncle' r) | _ -> failwith "Assertion error: invalid bourdoncle ist" -let build_bourdoncle'' (f : coq_function) : B.bourdoncle = +let build_bourdoncle'' (f : bb coq_function) : B.bourdoncle = let cfg = build_cfg f in match build_bourdoncle' (get_bourdoncle cfg) with | [] -> failwith "assertion error: empty bourdoncle" @@ -103,7 +102,7 @@ let build_order (b : B.bourdoncle) : coq_N PMap.t = let bo = build_order' (linearize b) (succ_pos N0) in (succ_pos N0, bo) -let build_bourdoncle (f : coq_function) : (B.bourdoncle * coq_N PMap.t) = +let build_bourdoncle (f : bb coq_function) : (B.bourdoncle * coq_N PMap.t) = let b = build_bourdoncle'' f in let bo = build_order b in (b, bo) diff --git a/src/extraction/Extraction.v b/src/extraction/Extraction.v index 9b82e1b..395ec47 100644 --- a/src/extraction/Extraction.v +++ b/src/extraction/Extraction.v @@ -24,9 +24,10 @@ From vericert Require RTLPar RTLBlockInstr HTLgen - Pipeline + (*Pipeline*) HLSOpts Predicate + Bourdoncle . From Coq Require DecidableClass. @@ -182,10 +183,13 @@ Extract Inlined Constant Binary.B2R => "fun _ -> assert false". Extract Inlined Constant Binary.round_mode => "fun _ -> assert false". Extract Inlined Constant Bracket.inbetween_loc => "fun _ -> assert false". -Extract Constant Pipeline.pipeline => "pipelining.pipeline". +(*Extract Constant Pipeline.pipeline => "pipelining.pipeline".*) Extract Constant RTLBlockgen.partition => "Partition.partition". Extract Constant RTLPargen.schedule => "Schedule.schedule_fn". +(* Loop normalization *) +Extract Constant IfConversion.build_bourdoncle => "BourdoncleAux.build_bourdoncle". + (* Needed in Coq 8.4 to avoid problems with Function definitions. *) Set Extraction AccessOpaque. @@ -195,9 +199,9 @@ Separate Extraction vericert.Compiler.transf_hls_temp RTLBlockgen.transl_program RTLBlockInstr.successors_instr HTLgen.tbl_to_case_expr - Pipeline.pipeline Predicate.sat_pred_simple Verilog.stmnt_to_list + Bourdoncle.bourdoncle Compiler.transf_c_program Compiler.transf_cminor_program Cexec.do_initial_state Cexec.do_step Cexec.at_final_state 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 26dcc7c33db422e93082c25d72de04fd4d1375ca Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 9 Dec 2021 20:06:43 +0000 Subject: Update the documentation --- docs | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/docs b/docs index 36abd86..be7d734 160000 --- a/docs +++ b/docs @@ -1 +1 @@ -Subproject commit 36abd86820f7521fcebb3b173acbcb6409b148b8 +Subproject commit be7d734e2cbf3f2619f83ad284cfd31f4fc36821 -- cgit From bf4a0ae5142e7838c322b6009ff5cdcbaf2dfd42 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 9 Dec 2021 20:07:40 +0000 Subject: Add PrintLoops for bourdoncle code --- src/bourdoncle/PrintLoops.ml | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) create mode 100644 src/bourdoncle/PrintLoops.ml diff --git a/src/bourdoncle/PrintLoops.ml b/src/bourdoncle/PrintLoops.ml new file mode 100644 index 0000000..f70f075 --- /dev/null +++ b/src/bourdoncle/PrintLoops.ml @@ -0,0 +1,22 @@ +open Camlcoq + +let indent i = String.make (2 * i) ' ' + +let rec print_bourdoncle i pp = + function + | Bourdoncle.I n -> + Printf.fprintf pp "%sI %d\n" (indent i) (P.to_int n) + | Bourdoncle.L (h, b) -> + Printf.fprintf pp "%sL %d\n" (indent i) (P.to_int h); + List.iter (fun a -> Printf.fprintf pp "%a" (print_bourdoncle (i+1)) a) b + +let print_if optdest loop = + match !optdest with + | None -> () + | Some f -> + let oc = open_out f in + print_bourdoncle 0 oc loop; + close_out oc + +let destination_loops : string option ref = ref None +let print_loops = print_if destination_loops -- cgit From a1d20be8f5001f6d833b530908b9ada25d140e1b Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 23 Feb 2022 15:44:05 +0000 Subject: Update gather data script --- scripts/gather_data.rkt | 49 ++++++++++++++++++++++++------------------------- 1 file changed, 24 insertions(+), 25 deletions(-) diff --git a/scripts/gather_data.rkt b/scripts/gather_data.rkt index 5e97c9e..c2da7d4 100644 --- a/scripts/gather_data.rkt +++ b/scripts/gather_data.rkt @@ -1,8 +1,13 @@ #lang racket +(require racket/match) +(require racket/list) +(require racket/file) + +(require threading) + (require xml) (require xml/path) -(require racket/match) (require csv-reading) (require csv-writing) @@ -11,25 +16,18 @@ (require data-frame) (require sawzall) -(require threading) - (permissive-xexprs #t) -(define (to-name f) (regexp-replace #rx".*/([^/]+)/encode_report.xml" f "\\1")) - -(define files - '("./data/data-mining/covariance/encode_report.xml" - "./data/stencils/heat-3d/encode_report.xml" - "./data/stencils/jacobi-1d/encode_report.xml" - "./data/stencils/seidel-2d/encode_report.xml" - "./data/stencils/jacobi-2d/encode_report.xml" - "./data/linear-algebra/kernels/doitgen/encode_report.xml" - "./data/linear-algebra/kernels/2mm/encode_report.xml" - "./data/linear-algebra/kernels/3mm/encode_report.xml" - "./data/linear-algebra/blas/gemver/encode_report.xml" - "./data/linear-algebra/blas/syrk/encode_report.xml" - "./data/linear-algebra/blas/gemm/encode_report.xml" - "./data/linear-algebra/solvers/trisolv/encode_report.xml")) +(define (to-name f) (regexp-replace #rx".*/([^/]+)" f "\\1")) + +(define (write-file file data) + (with-output-to-file file + (lambda () + (display data)) + #:exists 'replace)) + +(define files (file->lines "../benchmarks/polybench-syn/benchmark-list-master")) +(define base "/home/ymherklotz/projects/mpardalos-vericert/results-vericert-fun-full-inlining") (define (list->hash l) (foldr (lambda (v l) @@ -42,7 +40,7 @@ (hash) l)) (define name-f-map - (list->hash (map (lambda (f) (list (to-name f) f)) files))) + (list->hash (map (lambda (f) (list (to-name f) (string-append base "/" f "_report.xml"))) files))) (define (parse-vivado-report f) (let* ([encode-xml-port (open-input-file f)] @@ -58,6 +56,8 @@ (define (vivado-report-f f) (process-vivado-report (parse-vivado-report f))) ;;(vivado-report-f "./data/data-mining/covariance/encode_report.xml") +;; +;;(vivado-report-f "/home/ymherklotz/projects/mpardalos-vericert/results-vericert-fun-full-inlining/medley/nussinov_report.xml") (define synth-f (map flatten (hash-map name-f-map (lambda (n f) (let ([x (vivado-report-f f)]) @@ -78,13 +78,12 @@ (close-input-port exec-csv) report)) -(define sim-report (list->hashn (parse-sim-report "exec.csv"))) +(define sim-report (list->hashn (parse-sim-report (string-append base "/exec.csv")))) -(define csv (open-output-file "out2.csv")) -(display (table->string - (append '((benchmark lut_flip_flop slice regs luts ramfifo iopin dsps power delay cycles)) - (map (lambda (x) (append x (list (hash-ref sim-report (car x))))) synth-f))) csv) -(close-output-port csv) +(write-file "out.csv" + (table->string + (append '((benchmark lut_flip_flop slice regs luts ramfifo iopin dsps power delay cycles)) + (map (lambda (x) (append x (list (hash-ref sim-report (car x))))) synth-f)))) ;;(define vivado-report (process-vivado-report (parse-vivado-report "encode_report.xml"))) ;;(define cycles (hash-ref sim-report "covariance")) -- cgit From 9e35f6174b012d1ca7eef9921169e34c62c993dd Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 24 Feb 2022 15:16:24 +0000 Subject: Delete docs submodule --- .gitmodules | 3 --- docs | 1 - 2 files changed, 4 deletions(-) delete mode 160000 docs diff --git a/.gitmodules b/.gitmodules index 87bab2d..f30817e 100644 --- a/.gitmodules +++ b/.gitmodules @@ -1,6 +1,3 @@ [submodule "lib/CompCert"] path = lib/CompCert url = https://github.com/ymherklotz/CompCert.git -[submodule "docs"] - path = docs - url = https://github.com/ymherklotz/vericert-docs diff --git a/docs b/docs deleted file mode 160000 index be7d734..0000000 --- a/docs +++ /dev/null @@ -1 +0,0 @@ -Subproject commit be7d734e2cbf3f2619f83ad284cfd31f4fc36821 -- cgit From e043d3bd21e9a0ebd37b8ac2ca262ed630a5b192 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 25 Feb 2022 10:10:43 +0000 Subject: Add back pure documentation --- docs/README.org | 7 + docs/common.org | 13 ++ docs/documentation.org | 557 +++++++++++++++++++++++++++++++++++++++++++++++++ docs/fdl.org | 489 +++++++++++++++++++++++++++++++++++++++++++ docs/man.org | 36 ++++ 5 files changed, 1102 insertions(+) create mode 100644 docs/README.org create mode 100644 docs/common.org create mode 100644 docs/documentation.org create mode 100644 docs/fdl.org create mode 100644 docs/man.org diff --git a/docs/README.org b/docs/README.org new file mode 100644 index 0000000..4113ed9 --- /dev/null +++ b/docs/README.org @@ -0,0 +1,7 @@ +#+title: Vericert Documentation + +The documentation for [[https://github.com/ymherklotz/vericert][Vericert]], which is written in the [[/documentation.org][documentation.org]] file. + +* How to develop on documentation + +The documentation uses =hugo= to generate the website, diff --git a/docs/common.org b/docs/common.org new file mode 100644 index 0000000..8c39be0 --- /dev/null +++ b/docs/common.org @@ -0,0 +1,13 @@ +#+author: Yann Herklotz +#+email: git@ymhg.org + +#+macro: version 1.2.2 +#+macro: modified 2022-02-24 + +#+macro: link src_emacs-lisp[:results raw]{(ymhg/link "$1" "$2")} + +#+macro: texinfo_head (eval (if (eq org-export-current-backend 'texinfo) "#+exclude_tags: noexport_texinfo" "#+exclude_tags: onlyexport_texinfo")) +#+macro: latex_head (eval (if (eq org-export-current-backend 'latex) "#+exclude_tags: noexport_latex" "#+exclude_tags: onlyexport_latex")) +#+macro: hugo_head (eval (if (eq org-export-current-backend 'hugo) "#+exclude_tags: noexport_hugo" "#+exclude_tags: onlyexport_hugo")) +{{{texinfo_head}}} +{{{latex_head}}} diff --git a/docs/documentation.org b/docs/documentation.org new file mode 100644 index 0000000..d4ef799 --- /dev/null +++ b/docs/documentation.org @@ -0,0 +1,557 @@ +#+title: Vericert Manual +#+subtitle: Release {{{version}}} +#+author: Yann Herklotz +#+email: git@ymhg.org +#+language: en + +* Docs +:PROPERTIES: +:EXPORT_FILE_NAME: _index +:EXPORT_HUGO_SECTION: docs +:CUSTOM_ID: docs +:END: + +Vericert translates C code into a hardware description language called Verilog, which can then be +synthesised into hardware, to be placed onto a field-programmable gate array (FPGA) or +application-specific integrated circuit (ASIC). + +#+attr_html: :width 600 +#+caption: Current design of Vericert, where HTL is an intermediate language representing a finite state machine with data-path (FSMD) and Verilog is the target language. +#+name: fig:design +[[/images/toolflow.svg]] + +The design shown in Figure [[fig:design]] shows how Vericert leverages an existing verified C compiler +called [[https://compcert.org/compcert-C.html][CompCert]] to perform this translation. + +#+texinfo: @insertcopying + +* COPYING +:PROPERTIES: +:COPYING: t +:END: + +Copyright (C) 2019-2022 Yann Herklotz. + +#+begin_quote +Permission is granted to copy, distribute and/or modify this document +under the terms of the GNU Free Documentation License, Version 1.3 +or any later version published by the Free Software Foundation; +with no Invariant Sections, no Front-Cover Texts, and no Back-Cover +Texts. A copy of the license is included in the section entitled ``GNU +Free Documentation License''. +#+end_quote + +* Building Vericert +:PROPERTIES: +:EXPORT_FILE_NAME: building +:EXPORT_HUGO_SECTION: docs +:CUSTOM_ID: building +:END: + +#+transclude: [[file:~/projects/vericert/README.org::#building][file:~/projects/vericert/README.org::#building]] :only-contents :exclude-elements "headline property-drawer" +#+transclude: [[file:~/projects/vericert/README.org::#downloading-compcert][file:~/projects/vericert/README.org::#downloading-compcert]] +#+transclude: [[file:~/projects/vericert/README.org::#setting-up-nix][file:~/projects/vericert/README.org::#setting-up-nix]] +#+transclude: [[file:~/projects/vericert/README.org::#makefile-build][file:~/projects/vericert/README.org::#makefile-build]] + +** Testing + +To test out ~vericert~ you can try the following examples which are in the test folder using the +following: + +#+begin_src shell +./bin/vericert test/loop.c -o loop.v +./bin/vericert test/conditional.c -o conditional.v +./bin/vericert test/add.c -o add.v +#+end_src + +Or by running the test suite using the following command: + +#+begin_src shell +make test +#+end_src + +* Using Vericert +:PROPERTIES: +:CUSTOM_ID: using-vericert +:END: + +Vericert can be used to translate a subset of C into Verilog. As a simple example, consider the +following C file (~main.c~): + +#+begin_src C +void matrix_multiply(int first[2][2], int second[2][2], int multiply[2][2]) { + int sum = 0; + for (int c = 0; c < 2; c++) { + for (int d = 0; d < 2; d++) { + for (int k = 0; k < 2; k++) { + sum = sum + first[c][k]*second[k][d]; + } + multiply[c][d] = sum; + sum = 0; + } + } +} + +int main() { + int f[2][2] = {{1, 2}, {3, 4}}; + int s[2][2] = {{5, 6}, {7, 8}}; + int m[2][2] = {{0, 0}, {0, 0}}; + + matrix_multiply(f, s, m); + return m[1][1]; +} +#+end_src + +It can be compiled using the following command, assuming that vericert is somewhere on the path. + +#+begin_src shell +vericert main.c -o main.v +#+end_src + +The Verilog file contains a top-level test-bench, which can be given to any Verilog simulator to +simulate the hardware, which should give the same result as executing the C code. Using [[http://iverilog.icarus.com/][Icarus +Verilog]] as an example: + +#+begin_src shell +iverilog -o main_v main.v +#+end_src + +When executing, it should therefore print the following: + +#+begin_src shell +$ ./main_v +finished: 50 +#+end_src + +This gives the same result as executing the C in the following way: + +#+begin_src shell +$ gcc -o main_c main.c +$ ./main_c +$ echo $? +50 +#+end_src + +* Unreleased Features +:PROPERTIES: +:EXPORT_FILE_NAME: unreleased +:EXPORT_HUGO_SECTION: docs +:CUSTOM_ID: unreleased-features +:END: + +The following are unreleased features in Vericert that are currently being worked on and have not +been completely proven correct yet. Currently this includes features such as: + +- [[#scheduling][scheduling]], +- [[#operation-chaining][operation chaining]], +- [[#if-conversion][if-conversion]], and +- [[#functions][functions]]. + +This page gives some preliminary information on how the features are implemented and how the proofs +for the features are being done. Once these features are properly implemented, they will be added +to the proper documentation. + +** Scheduling +:PROPERTIES: +:CUSTOM_ID: scheduling +:END: + +Scheduling is an optimisation which is used to run various instructions in parallel that are +independent to each other. + +** Operation Chaining +:PROPERTIES: +:CUSTOM_ID: operation-chaining +:END: + +Operation chaining is an optimisation that can be added on to scheduling and allows for the +sequential execution of instructions in a clock cycle, while executing other instructions in +parallel in the same clock cycle. + +** If-conversion +:PROPERTIES: +:CUSTOM_ID: if-conversion +:END: + +If-conversion is an optimisation which can turn code with simple control flow into a single block +(called a hyper-block), using predicated instructions. + +** Functions +:PROPERTIES: +:CUSTOM_ID: functions +:END: + +Functions are currently only inlined in Vericert, however, we are working on a proper interface to +integrate function calls into the hardware. + +* Coq Style Guide + :PROPERTIES: + :CUSTOM_ID: coq-style-guide + :EXPORT_FILE_NAME: coq-style-guide + :EXPORT_HUGO_SECTION: docs + :END: + +This style guide was taken from [[https://github.com/project-oak/silveroak][Silveroak]], it outlines code style for Coq code in this +repository. There are certainly other valid strategies and opinions on Coq code style; this is laid +out purely in the name of consistency. For a visual example of the style, see the [[#example][example]] at the +bottom of this file. + +** Code organization + :PROPERTIES: + :CUSTOM_ID: code-organization + :END: +*** Legal banner + :PROPERTIES: + :CUSTOM_ID: legal-banner + :END: + +- Files should begin with a copyright/license banner, as shown in the example above. + +*** Import statements + :PROPERTIES: + :CUSTOM_ID: import-statements + :END: + +- =Require Import= statements should all go at the top of the file, followed by file-wide =Import= + statements. + + - =Import=s often contain notations or typeclass instances that might override notations or + instances from another library, so it's nice to highlight them separately. + +- One =Require Import= statement per line; it's easier to scan that way. +- =Require Import= statements should use "fully-qualified" names (e.g. =Require Import + Coq.ZArith.ZArith= instead of =Require Import ZArith=). + + - Use the =Locate= command to find the fully-qualified name! + +- =Require Import='s should go in the following order: + + 1. Standard library dependencies (start with =Coq.=) + 2. External dependencies (anything outside the current project) + 3. Same-project dependencies + +- =Require Import='s with the same root library (the name before the first =.=) should be grouped + together. Within each root-library group, they should be in alphabetical order (so =Coq.Lists.List= + before =Coq.ZArith.ZArith=). + +*** Notations and scopes + :PROPERTIES: + :CUSTOM_ID: notations-and-scopes + :END: + +- Any file-wide =Local Open Scope='s should come immediately after the =Import=s (see example). + + - Always use =Local Open Scope=; just =Open Scope= will sneakily open the scope for those who import + your file. + +- Put notations in their own separate modules or files, so that those who import your file can + choose whether or not they want the notations. + + - Conflicting notations can cause a lot of headache, so it comes in very handy to leave this + flexibility! + +** Formatting + :PROPERTIES: + :CUSTOM_ID: formatting + :END: +*** Line length + :PROPERTIES: + :CUSTOM_ID: line-length + :END: + +- Maximum line length 80 characters. + + - Many Coq IDE setups divide the screen in half vertically and use only half to display source + code, so more than 80 characters can be genuinely hard to read on a laptop. + +*** Whitespace and indentation + :PROPERTIES: + :CUSTOM_ID: whitespace-and-indentation + :END: + +- No trailing whitespace. + +- Spaces, not tabs. + +- Files should end with a newline. + + - Many editors do this automatically on save. + +- Colons may be either "English-spaced", with no space before the colon and one space after (=x: nat=) + or "French-spaced", with one space before and after (=x : nat=). + +- Default indentation is 2 spaces. + + - Keeping this small prevents complex proofs from being indented ridiculously far, and matches IDE + defaults. + +- Use 2-space indents if inserting a line break immediately after: + + - =Proof.= + - =fun <...> =>= + - =forall <...>,= + - =exists <....>,= + +- The style for indenting arguments in function application depends on where you make a line + break. If you make the line break immediately after the function name, use a 2-space + indent. However, if you make it after one or more arguments, align the next line with the first + argument: + + #+begin_src coq + (Z.pow + 1 2) + (Z.pow 1 2 3 + 4 5 6) + #+end_src + +- =Inductive= cases should not be indented. Example: + + #+begin_src coq + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo + . + #+end_src + +- =match= or =lazymatch= cases should line up with the "m" in =match= or "l" in =lazymatch=, as in the + following examples: + + #+begin_src coq + match x with + | 3 => true + | _ => false + end. + + lazymatch x with + | 3 => idtac + | _ => fail "Not equal to 3:" x + end. + + repeat match goal with + | _ => progress subst + | _ => reflexivity + end. + + do 2 lazymatch goal with + | |- context [eq] => idtac + end. + #+end_src + +** Definitions and Fixpoints + :PROPERTIES: + :CUSTOM_ID: definitions-and-fixpoints + :END: + +- It's okay to leave the return type of =Definition='s and =Fixpoint='s implicit (e.g. ~Definition x := 5~ + instead of ~Definition x : nat := 5~) when the type is very simple or obvious (for instance, the + definition is in a file which deals exclusively with operations on =Z=). + +** Inductives + :PROPERTIES: + :CUSTOM_ID: inductives + :END: + +- The =.= ending an =Inductive= can be either on the same line as the last case or on its own line + immediately below. That is, both of the following are acceptable: + + #+begin_src coq + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo + . + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo. + #+end_src + +** Lemma/Theorem statements + :PROPERTIES: + :CUSTOM_ID: lemmatheorem-statements + :END: + +- Generally, use =Theorem= for the most important, top-level facts you prove and =Lemma= for everything + else. +- Insert a line break after the colon in the lemma statement. +- Insert a line break after the comma for =forall= or =exist= quantifiers. +- Implication arrows (=->=) should share a line with the previous hypothesis, not the following one. +- There is no need to make a line break after every =->=; short preconditions may share a line. + +** Proofs and tactics + :PROPERTIES: + :CUSTOM_ID: proofs-and-tactics + :END: + +- Use the =Proof= command (lined up vertically with =Lemma= or =Theorem= it corresponds to) to open a + proof, and indent the first line after it 2 spaces. + +- Very small proofs (where =Proof. Qed.= is <= 80 characters) can go all in one line. + +- When ending a proof, align the ending statement (=Qed=, =Admitted=, etc.) with =Proof=. + +- Avoid referring to autogenerated names (e.g. =H0=, =n0=). It's okay to let Coq generate these names, + but you should not explicitly refer to them in your proof. So =intros; my_solver= is fine, but + =intros; apply H1; my_solver= is not fine. + + - You can force a non-autogenerated name by either putting the variable before the colon in the + lemma statement (=Lemma foo x : ...= instead of =Lemma foo : forall x, ...=), or by passing + arguments to =intros= (e.g. =intros ? x= to name the second argument =x=) + +- This way, the proof won't break when new hypotheses are added or autogenerated variable names + change. + +- Use curly braces ={}= for subgoals, instead of bullets. + +- /Never write tactics with more than one subgoal focused./ This can make the proof very confusing to + step through! If you have more than one subgoal, use curly braces. + +- Consider adding a comment after the opening curly brace that explains what case you're in (see + example). + + - This is not necessary for small subgoals but can help show the major lines of reasoning in large + proofs. + +- If invoking a tactic that is expected to return multiple subgoals, use =[ | ... | ]= before the =.= to + explicitly specify how many subgoals you expect. + + - Examples: =split; [ | ].= =induction z; [ | | ].= + - This helps make code more maintainable, because it fails immediately if your tactic no longer + solves as many subgoals as expected (or unexpectedly solves more). + +- If invoking a string of tactics (composed by =;=) that will break the goal into multiple subgoals + and then solve all but one, still use =[ ]= to enforce that all but one goal is solved. + + - Example: =split; try lia; [ ]=. + +- Tactics that consist only of =repeat=-ing a procedure (e.g. =repeat match=, =repeat first=) should + factor out a single step of that procedure a separate tactic called =_step=, because + the single-step version is much easier to debug. For instance: + + #+begin_src coq + Ltac crush_step := + match goal with + | _ => progress subst + | _ => reflexivity + end. + Ltac crush := repeat crush_step. + #+end_src + +** Naming + :PROPERTIES: + :CUSTOM_ID: naming + :END: + +- Helper proofs about standard library datatypes should go in a module that is named to match the + standard library module (see example). + + - This makes the helper proofs look like standard-library ones, which is helpful for categorizing + them if they're genuinely at the standard-library level of abstraction. + +- Names of modules should start with capital letters. +- Names of inductives and their constructors should start with capital letters. +- Names of other definitions/lemmas should be snake case. + +** Example + :PROPERTIES: + :CUSTOM_ID: example + :END: +A small standalone Coq file that exhibits many of the style points. + +#+begin_src coq +(* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2021 Name + * + * + *) + + Require Import Coq.Lists.List. + Require Import Coq.micromega.Lia. + Require Import Coq.ZArith.ZArith. + Import ListNotations. + Local Open Scope Z_scope. + + (* Helper proofs about standard library integers (Z) go within [Module Z] so + that they match standard-library Z lemmas when used. *) + Module Z. + Lemma pow_3_r x : x ^ 3 = x * x * x. + Proof. lia. Qed. (* very short proofs can go all on one line *) + + Lemma pow_4_r x : x ^ 4 = x * x * x * x. + Proof. + change 4 with (Z.succ (Z.succ (Z.succ (Z.succ 0)))). + repeat match goal with + | _ => rewrite Z.pow_1_r + | _ => rewrite Z.pow_succ_r by lia + | |- context [x * (?a * ?b)] => + replace (x * (a * b)) with (a * b * x) by lia + | _ => reflexivity + end. + Qed. + End Z. + (* Now we can access the lemmas above as Z.pow_3_r and Z.pow_4_r, as if they + were in the ZArith library! *) + + Definition bar (x y : Z) := x ^ (y + 1). + + (* example with a painfully manual proof to show case formatting *) + Lemma bar_upper_bound : + forall x y a, + 0 <= x <= a -> 0 <= y -> + 0 <= bar x y <= a ^ (y + 1). + Proof. + (* avoid referencing autogenerated names by explicitly naming variables *) + intros x y a Hx Hy. revert y Hy x a Hx. + (* explicitly indicate # subgoals with [ | ... | ] if > 1 *) + cbv [bar]; refine (natlike_ind _ _ _); [ | ]. + { (* y = 0 *) + intros; lia. } + { (* y = Z.succ _ *) + intros. + rewrite Z.add_succ_l, Z.pow_succ_r by lia. + split. + { (* 0 <= bar x y *) + apply Z.mul_nonneg_nonneg; [ lia | ]. + apply Z.pow_nonneg; lia. } + { (* bar x y < a ^ y *) + rewrite Z.pow_succ_r by lia. + apply Z.mul_le_mono_nonneg; try lia; + [ apply Z.pow_nonneg; lia | ]. + (* For more flexible proofs, use match statements to find hypotheses + rather than referring to them by autogenerated names like H0. In this + case, we'll take any hypothesis that applies to and then solves the + goal. *) + match goal with H : _ |- _ => apply H; solve [auto] end. } } + Qed. + + (* Put notations in a separate module or file so that importers can + decide whether or not to use them. *) + Module BarNotations. + Infix "#" := bar (at level 40) : Z_scope. + Notation "x '##'" := (bar x x) (at level 40) : Z_scope. + End BarNotations. +#+end_src + +* Index - Features +:PROPERTIES: +:CUSTOM_ID: cindex +:APPENDIX: t +:INDEX: cp +:DESCRIPTION: Key concepts & features +:END: + +* Export Setup :noexport: + +#+setupfile: common.org + +#+export_file_name: vericert.texi + +#+texinfo_dir_category: High-level synthesis tool +#+texinfo_dir_title: Vericert +#+texinfo_dir_desc: Formally verified high-level synthesis tool + +* GNU Free Documentation License +:PROPERTIES: +:appendix: t +:END: + +#+include: fdl.org diff --git a/docs/fdl.org b/docs/fdl.org new file mode 100644 index 0000000..81e2cd9 --- /dev/null +++ b/docs/fdl.org @@ -0,0 +1,489 @@ +# The GNU Free Documentation License. +#+begin_center +Version 1.3, 3 November 2008 +#+end_center + +# This file is intended to be included within another document. + +#+begin_verse +Copyright \copy{} 2000, 2001, 2002, 2007, 2008 Free Software Foundation, Inc. +https://fsf.org/ + +Everyone is permitted to copy and distribute verbatim copies +of this license document, but changing it is not allowed. +#+end_verse + +0. [@0] PREAMBLE + + The purpose of this License is to make a manual, textbook, or other + functional and useful document @@texinfo:@dfn{@@free@@texinfo:}@@ + in the sense of freedom: to assure everyone the effective freedom + to copy and redistribute it, with or without modifying it, either + commercially or noncommercially. Secondarily, this License + preserves for the author and publisher a way to get credit for + their work, while not being considered responsible for + modifications made by others. + + This License is a kind of "copyleft", which means that derivative + works of the document must themselves be free in the same sense. + It complements the GNU General Public License, which is a copyleft + license designed for free software. + + We have designed this License in order to use it for manuals for + free software, because free software needs free documentation: + a free program should come with manuals providing the same freedoms + that the software does. But this License is not limited to + software manuals; it can be used for any textual work, regardless + of subject matter or whether it is published as a printed book. We + recommend this License principally for works whose purpose is + instruction or reference. + +1. APPLICABILITY AND DEFINITIONS + + This License applies to any manual or other work, in any medium, + that contains a notice placed by the copyright holder saying it can + be distributed under the terms of this License. Such a notice + grants a world-wide, royalty-free license, unlimited in duration, + to use that work under the conditions stated herein. The + "Document", below, refers to any such manual or work. Any member + of the public is a licensee, and is addressed as "you". You accept + the license if you copy, modify or distribute the work in a way + requiring permission under copyright law. + + A "Modified Version" of the Document means any work containing the + Document or a portion of it, either copied verbatim, or with + modifications and/or translated into another language. + + A "Secondary Section" is a named appendix or a front-matter section + of the Document that deals exclusively with the relationship of the + publishers or authors of the Document to the Document's overall + subject (or to related matters) and contains nothing that could + fall directly within that overall subject. (Thus, if the Document + is in part a textbook of mathematics, a Secondary Section may not + explain any mathematics.) The relationship could be a matter of + historical connection with the subject or with related matters, or + of legal, commercial, philosophical, ethical or political position + regarding them. + + The "Invariant Sections" are certain Secondary Sections whose + titles are designated, as being those of Invariant Sections, in the + notice that says that the Document is released under this License. + If a section does not fit the above definition of Secondary then it + is not allowed to be designated as Invariant. The Document may + contain zero Invariant Sections. If the Document does not identify + any Invariant Sections then there are none. + + The "Cover Texts" are certain short passages of text that are + listed, as Front-Cover Texts or Back-Cover Texts, in the notice + that says that the Document is released under this License. + A Front-Cover Text may be at most 5 words, and a Back-Cover Text + may be at most 25 words. + + A "Transparent" copy of the Document means a machine-readable copy, + represented in a format whose specification is available to the + general public, that is suitable for revising the document + straightforwardly with generic text editors or (for images composed + of pixels) generic paint programs or (for drawings) some widely + available drawing editor, and that is suitable for input to text + formatters or for automatic translation to a variety of formats + suitable for input to text formatters. A copy made in an otherwise + Transparent file format whose markup, or absence of markup, has + been arranged to thwart or discourage subsequent modification by + readers is not Transparent. An image format is not Transparent if + used for any substantial amount of text. A copy that is not + "Transparent" is called "Opaque". + + Examples of suitable formats for Transparent copies include plain + ASCII without markup, Texinfo input format, LaTeX input format, + SGML or XML using a publicly available DTD, and standard-conforming + simple HTML, PostScript or PDF designed for human modification. + Examples of transparent image formats include PNG, XCF and JPG. + Opaque formats include proprietary formats that can be read and + edited only by proprietary word processors, SGML or XML for which + the DTD and/or processing tools are not generally available, and + the machine-generated HTML, PostScript or PDF produced by some word + processors for output purposes only. + + The "Title Page" means, for a printed book, the title page itself, + plus such following pages as are needed to hold, legibly, the + material this License requires to appear in the title page. For + works in formats which do not have any title page as such, "Title + Page" means the text near the most prominent appearance of the + work's title, preceding the beginning of the body of the text. + + The "publisher" means any person or entity that distributes copies + of the Document to the public. + + A section "Entitled XYZ" means a named subunit of the Document + whose title either is precisely XYZ or contains XYZ in parentheses + following text that translates XYZ in another language. (Here XYZ + stands for a specific section name mentioned below, such as + "Acknowledgements", "Dedications", "Endorsements", or "History".) + To "Preserve the Title" of such a section when you modify the + Document means that it remains a section "Entitled XYZ" according + to this definition. + + The Document may include Warranty Disclaimers next to the notice + which states that this License applies to the Document. These + Warranty Disclaimers are considered to be included by reference in + this License, but only as regards disclaiming warranties: any other + implication that these Warranty Disclaimers may have is void and + has no effect on the meaning of this License. + +2. VERBATIM COPYING + + You may copy and distribute the Document in any medium, either + commercially or noncommercially, provided that this License, the + copyright notices, and the license notice saying this License + applies to the Document are reproduced in all copies, and that you + add no other conditions whatsoever to those of this License. You + may not use technical measures to obstruct or control the reading + or further copying of the copies you make or distribute. However, + you may accept compensation in exchange for copies. If you + distribute a large enough number of copies you must also follow the + conditions in section 3. + + You may also lend copies, under the same conditions stated above, + and you may publicly display copies. + +3. COPYING IN QUANTITY + + If you publish printed copies (or copies in media that commonly + have printed covers) of the Document, numbering more than 100, and + the Document's license notice requires Cover Texts, you must + enclose the copies in covers that carry, clearly and legibly, all + these Cover Texts: Front-Cover Texts on the front cover, and + Back-Cover Texts on the back cover. Both covers must also clearly + and legibly identify you as the publisher of these copies. The + front cover must present the full title with all words of the title + equally prominent and visible. You may add other material on the + covers in addition. Copying with changes limited to the covers, as + long as they preserve the title of the Document and satisfy these + conditions, can be treated as verbatim copying in other respects. + + If the required texts for either cover are too voluminous to fit + legibly, you should put the first ones listed (as many as fit + reasonably) on the actual cover, and continue the rest onto + adjacent pages. + + If you publish or distribute Opaque copies of the Document + numbering more than 100, you must either include a machine-readable + Transparent copy along with each Opaque copy, or state in or with + each Opaque copy a computer-network location from which the general + network-using public has access to download using public-standard + network protocols a complete Transparent copy of the Document, free + of added material. If you use the latter option, you must take + reasonably prudent steps, when you begin distribution of Opaque + copies in quantity, to ensure that this Transparent copy will + remain thus accessible at the stated location until at least one + year after the last time you distribute an Opaque copy (directly or + through your agents or retailers) of that edition to the public. + + It is requested, but not required, that you contact the authors of + the Document well before redistributing any large number of copies, + to give them a chance to provide you with an updated version of the + Document. + +4. MODIFICATIONS + + You may copy and distribute a Modified Version of the Document + under the conditions of sections 2 and 3 above, provided that you + release the Modified Version under precisely this License, with the + Modified Version filling the role of the Document, thus licensing + distribution and modification of the Modified Version to whoever + possesses a copy of it. In addition, you must do these things in + the Modified Version: + + #+attr_texinfo: :enum A + 1. Use in the Title Page (and on the covers, if any) a title + distinct from that of the Document, and from those of previous + versions (which should, if there were any, be listed in the + History section of the Document). You may use the same title as + a previous version if the original publisher of that version + gives permission. + + 2. List on the Title Page, as authors, one or more persons or + entities responsible for authorship of the modifications in the + Modified Version, together with at least five of the principal + authors of the Document (all of its principal authors, if it has + fewer than five), unless they release you from this requirement. + + 3. State on the Title page the name of the publisher of the + Modified Version, as the publisher. + + 4. Preserve all the copyright notices of the Document. + + 5. Add an appropriate copyright notice for your modifications + adjacent to the other copyright notices. + + 6. Include, immediately after the copyright notices, a license + notice giving the public permission to use the Modified Version + under the terms of this License, in the form shown in the + Addendum below. + + 7. Preserve in that license notice the full lists of Invariant + Sections and required Cover Texts given in the Document's + license notice. + + 8. Include an unaltered copy of this License. + + 9. Preserve the section Entitled "History", Preserve its Title, and + add to it an item stating at least the title, year, new authors, + and publisher of the Modified Version as given on the Title + Page. If there is no section Entitled "History" in the Document, + create one stating the title, year, authors, and publisher of + the Document as given on its Title Page, then add an item + describing the Modified Version as stated in the previous + sentence. + + 10. Preserve the network location, if any, given in the Document + for public access to a Transparent copy of the Document, and + likewise the network locations given in the Document for + previous versions it was based on. These may be placed in the + "History" section. You may omit a network location for a work + that was published at least four years before the Document + itself, or if the original publisher of the version it refers + to gives permission. + + 11. For any section Entitled "Acknowledgements" or "Dedications", + Preserve the Title of the section, and preserve in the section + all the substance and tone of each of the contributor + acknowledgements and/or dedications given therein. + + 12. Preserve all the Invariant Sections of the Document, unaltered + in their text and in their titles. Section numbers or the + equivalent are not considered part of the section titles. + + 13. Delete any section Entitled "Endorsements". Such a section may + not be included in the Modified Version. + + 14. Do not retitle any existing section to be Entitled + "Endorsements" or to conflict in title with any Invariant + Section. + + 15. Preserve any Warranty Disclaimers. + + If the Modified Version includes new front-matter sections or + appendices that qualify as Secondary Sections and contain no material + copied from the Document, you may at your option designate some or all + of these sections as invariant. To do this, add their titles to the + list of Invariant Sections in the Modified Version's license notice. + These titles must be distinct from any other section titles. + + You may add a section Entitled "Endorsements", provided it contains + nothing but endorsements of your Modified Version by various + parties---for example, statements of peer review or that the text has + been approved by an organization as the authoritative definition of a + standard. + + You may add a passage of up to five words as a Front-Cover Text, and a + passage of up to 25 words as a Back-Cover Text, to the end of the list + of Cover Texts in the Modified Version. Only one passage of + Front-Cover Text and one of Back-Cover Text may be added by (or + through arrangements made by) any one entity. If the Document already + includes a cover text for the same cover, previously added by you or + by arrangement made by the same entity you are acting on behalf of, + you may not add another; but you may replace the old one, on explicit + permission from the previous publisher that added the old one. + + The author(s) and publisher(s) of the Document do not by this License + give permission to use their names for publicity for or to assert or + imply endorsement of any Modified Version. + +5. COMBINING DOCUMENTS + + You may combine the Document with other documents released under + this License, under the terms defined in section 4 above for + modified versions, provided that you include in the combination all + of the Invariant Sections of all of the original documents, + unmodified, and list them all as Invariant Sections of your + combined work in its license notice, and that you preserve all + their Warranty Disclaimers. + + The combined work need only contain one copy of this License, and + multiple identical Invariant Sections may be replaced with a single + copy. If there are multiple Invariant Sections with the same name + but different contents, make the title of each such section unique + by adding at the end of it, in parentheses, the name of the + original author or publisher of that section if known, or else + a unique number. Make the same adjustment to the section titles in + the list of Invariant Sections in the license notice of the + combined work. + + In the combination, you must combine any sections Entitled + "History" in the various original documents, forming one section + Entitled "History"; likewise combine any sections Entitled + "Acknowledgements", and any sections Entitled "Dedications". You + must delete all sections Entitled "Endorsements." + +6. COLLECTIONS OF DOCUMENTS + + You may make a collection consisting of the Document and other + documents released under this License, and replace the individual + copies of this License in the various documents with a single copy + that is included in the collection, provided that you follow the + rules of this License for verbatim copying of each of the documents + in all other respects. + + You may extract a single document from such a collection, and + distribute it individually under this License, provided you insert + a copy of this License into the extracted document, and follow this + License in all other respects regarding verbatim copying of that + document. + +7. AGGREGATION WITH INDEPENDENT WORKS + + A compilation of the Document or its derivatives with other + separate and independent documents or works, in or on a volume of + a storage or distribution medium, is called an "aggregate" if the + copyright resulting from the compilation is not used to limit the + legal rights of the compilation's users beyond what the individual + works permit. When the Document is included in an aggregate, this + License does not apply to the other works in the aggregate which + are not themselves derivative works of the Document. + + If the Cover Text requirement of section 3 is applicable to these + copies of the Document, then if the Document is less than one half + of the entire aggregate, the Document's Cover Texts may be placed + on covers that bracket the Document within the aggregate, or the + electronic equivalent of covers if the Document is in electronic + form. Otherwise they must appear on printed covers that bracket + the whole aggregate. + +8. TRANSLATION + + Translation is considered a kind of modification, so you may + distribute translations of the Document under the terms of + section 4. Replacing Invariant Sections with translations requires + special permission from their copyright holders, but you may + include translations of some or all Invariant Sections in addition + to the original versions of these Invariant Sections. You may + include a translation of this License, and all the license notices + in the Document, and any Warranty Disclaimers, provided that you + also include the original English version of this License and the + original versions of those notices and disclaimers. In case of + a disagreement between the translation and the original version of + this License or a notice or disclaimer, the original version will + prevail. + + If a section in the Document is Entitled "Acknowledgements", + "Dedications", or "History", the requirement (section 4) to + Preserve its Title (section 1) will typically require changing the + actual title. + +9. TERMINATION + + You may not copy, modify, sublicense, or distribute the Document + except as expressly provided under this License. Any attempt + otherwise to copy, modify, sublicense, or distribute it is void, + and will automatically terminate your rights under this License. + + However, if you cease all violation of this License, then your + license from a particular copyright holder is reinstated (a) + provisionally, unless and until the copyright holder explicitly and + finally terminates your license, and (b) permanently, if the + copyright holder fails to notify you of the violation by some + reasonable means prior to 60 days after the cessation. + + Moreover, your license from a particular copyright holder is + reinstated permanently if the copyright holder notifies you of the + violation by some reasonable means, this is the first time you have + received notice of violation of this License (for any work) from + that copyright holder, and you cure the violation prior to 30 days + after your receipt of the notice. + + Termination of your rights under this section does not terminate + the licenses of parties who have received copies or rights from you + under this License. If your rights have been terminated and not + permanently reinstated, receipt of a copy of some or all of the + same material does not give you any rights to use it. + +10. FUTURE REVISIONS OF THIS LICENSE + + The Free Software Foundation may publish new, revised versions of + the GNU Free Documentation License from time to time. Such new + versions will be similar in spirit to the present version, but may + differ in detail to address new problems or concerns. See + https://www.gnu.org/copyleft/. + + Each version of the License is given a distinguishing version + number. If the Document specifies that a particular numbered + version of this License "or any later version" applies to it, you + have the option of following the terms and conditions either of + that specified version or of any later version that has been + published (not as a draft) by the Free Software Foundation. If + the Document does not specify a version number of this License, + you may choose any version ever published (not as a draft) by the + Free Software Foundation. If the Document specifies that a proxy + can decide which future versions of this License can be used, that + proxy's public statement of acceptance of a version permanently + authorizes you to choose that version for the Document. + +11. RELICENSING + + "Massive Multiauthor Collaboration Site" (or "MMC Site") means any + World Wide Web server that publishes copyrightable works and also + provides prominent facilities for anybody to edit those works. + A public wiki that anybody can edit is an example of such + a server. A "Massive Multiauthor Collaboration" (or "MMC") + contained in the site means any set of copyrightable works thus + published on the MMC site. + + "CC-BY-SA" means the Creative Commons Attribution-Share Alike 3.0 + license published by Creative Commons Corporation, + a not-for-profit corporation with a principal place of business in + San Francisco, California, as well as future copyleft versions of + that license published by that same organization. + + "Incorporate" means to publish or republish a Document, in whole + or in part, as part of another Document. + + An MMC is "eligible for relicensing" if it is licensed under this + License, and if all works that were first published under this + License somewhere other than this MMC, and subsequently + incorporated in whole or in part into the MMC, (1) had no cover + texts or invariant sections, and (2) were thus incorporated prior + to November 1, 2008. + + The operator of an MMC Site may republish an MMC contained in the + site under CC-BY-SA on the same site at any time before August 1, + 2009, provided the MMC is eligible for relicensing. + +#+texinfo: @page + +* ADDENDUM: How to use this License for your documents +:PROPERTIES: +:UNNUMBERED: notoc +:END: + +To use this License in a document you have written, include a copy of +the License in the document and put the following copyright and +license notices just after the title page: + +#+begin_example + Copyright (C) YEAR YOUR NAME. + Permission is granted to copy, distribute and/or modify this document + under the terms of the GNU Free Documentation License, Version 1.3 + or any later version published by the Free Software Foundation; + with no Invariant Sections, no Front-Cover Texts, and no Back-Cover + Texts. A copy of the license is included in the section entitled ``GNU + Free Documentation License''. +#+end_example + +If you have Invariant Sections, Front-Cover Texts and Back-Cover Texts, +replace the "with...Texts."\nbsp{}line with this: + +#+begin_example + with the Invariant Sections being LIST THEIR TITLES, with + the Front-Cover Texts being LIST, and with the Back-Cover Texts + being LIST. +#+end_example + +If you have Invariant Sections without Cover Texts, or some other +combination of the three, merge those two alternatives to suit the +situation. + +If your document contains nontrivial examples of program code, we +recommend releasing these examples in parallel under your choice of +free software license, such as the GNU General Public License, to +permit their use in free software. diff --git a/docs/man.org b/docs/man.org new file mode 100644 index 0000000..279165c --- /dev/null +++ b/docs/man.org @@ -0,0 +1,36 @@ +#+title: vericert +#+man_class_options: :section-id "1" +#+export_file_name: vericert.1 + +* NAME + +vericert - A formally verified high-level synthesis tool. + +* SYNOPSYS + +*vericert* [ *OPTION* ]... [ *FILE* ]... + +* DESCRIPTION + +- -drtl :: Print intermediate RTL. + +* AUTHOR + +Written by Yann Herklotz, Michalis Pardalos, James Pollard, Nadesh Ramanathan and John Wickerson. + +* COPYRIGHT + +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 . -- cgit From c1c5fc8e12342a9fe435c8066c8e9316036ff991 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 25 Feb 2022 13:39:17 +0000 Subject: Add more documentation and add coqdoc stylesheet --- Makefile | 4 + docs/common.org | 2 + docs/documentation.org | 4 +- docs/fdl.org | 489 ---------------------------- docs/man.org | 54 ++- docs/res/coqdoc.css | 867 +++++++++++++++++++++++++++++++++++++++++++++++++ docs/res/fdl.org | 489 ++++++++++++++++++++++++++++ 7 files changed, 1417 insertions(+), 492 deletions(-) delete mode 100644 docs/fdl.org create mode 100644 docs/res/coqdoc.css create mode 100644 docs/res/fdl.org diff --git a/Makefile b/Makefile index 5577ba9..9c40941 100644 --- a/Makefile +++ b/Makefile @@ -50,6 +50,10 @@ proof: Makefile.coq doc: Makefile.coq $(MAKE) COQDOCFLAGS="$(COQDOCFLAGS)" -f Makefile.coq html + cp ./docs/res/coqdoc.css html/. + +doc-pdf: Makefile.coq + $(MAKE) COQDOCFLAGS="$(COQDOCFLAGS)" -f Makefile.coq all.tex extraction: src/extraction/STAMP diff --git a/docs/common.org b/docs/common.org index 8c39be0..aa27264 100644 --- a/docs/common.org +++ b/docs/common.org @@ -4,6 +4,8 @@ #+macro: version 1.2.2 #+macro: modified 2022-02-24 +#+macro: base_url https://vericert.ymhg.org + #+macro: link src_emacs-lisp[:results raw]{(ymhg/link "$1" "$2")} #+macro: texinfo_head (eval (if (eq org-export-current-backend 'texinfo) "#+exclude_tags: noexport_texinfo" "#+exclude_tags: onlyexport_texinfo")) diff --git a/docs/documentation.org b/docs/documentation.org index d4ef799..3a0633f 100644 --- a/docs/documentation.org +++ b/docs/documentation.org @@ -18,7 +18,7 @@ application-specific integrated circuit (ASIC). #+attr_html: :width 600 #+caption: Current design of Vericert, where HTL is an intermediate language representing a finite state machine with data-path (FSMD) and Verilog is the target language. #+name: fig:design -[[/images/toolflow.svg]] +[[./images/toolflow.svg]] The design shown in Figure [[fig:design]] shows how Vericert leverages an existing verified C compiler called [[https://compcert.org/compcert-C.html][CompCert]] to perform this translation. @@ -554,4 +554,4 @@ A small standalone Coq file that exhibits many of the style points. :appendix: t :END: -#+include: fdl.org +#+include: res/fdl.org diff --git a/docs/fdl.org b/docs/fdl.org deleted file mode 100644 index 81e2cd9..0000000 --- a/docs/fdl.org +++ /dev/null @@ -1,489 +0,0 @@ -# The GNU Free Documentation License. -#+begin_center -Version 1.3, 3 November 2008 -#+end_center - -# This file is intended to be included within another document. - -#+begin_verse -Copyright \copy{} 2000, 2001, 2002, 2007, 2008 Free Software Foundation, Inc. -https://fsf.org/ - -Everyone is permitted to copy and distribute verbatim copies -of this license document, but changing it is not allowed. -#+end_verse - -0. [@0] PREAMBLE - - The purpose of this License is to make a manual, textbook, or other - functional and useful document @@texinfo:@dfn{@@free@@texinfo:}@@ - in the sense of freedom: to assure everyone the effective freedom - to copy and redistribute it, with or without modifying it, either - commercially or noncommercially. Secondarily, this License - preserves for the author and publisher a way to get credit for - their work, while not being considered responsible for - modifications made by others. - - This License is a kind of "copyleft", which means that derivative - works of the document must themselves be free in the same sense. - It complements the GNU General Public License, which is a copyleft - license designed for free software. - - We have designed this License in order to use it for manuals for - free software, because free software needs free documentation: - a free program should come with manuals providing the same freedoms - that the software does. But this License is not limited to - software manuals; it can be used for any textual work, regardless - of subject matter or whether it is published as a printed book. We - recommend this License principally for works whose purpose is - instruction or reference. - -1. APPLICABILITY AND DEFINITIONS - - This License applies to any manual or other work, in any medium, - that contains a notice placed by the copyright holder saying it can - be distributed under the terms of this License. Such a notice - grants a world-wide, royalty-free license, unlimited in duration, - to use that work under the conditions stated herein. The - "Document", below, refers to any such manual or work. Any member - of the public is a licensee, and is addressed as "you". You accept - the license if you copy, modify or distribute the work in a way - requiring permission under copyright law. - - A "Modified Version" of the Document means any work containing the - Document or a portion of it, either copied verbatim, or with - modifications and/or translated into another language. - - A "Secondary Section" is a named appendix or a front-matter section - of the Document that deals exclusively with the relationship of the - publishers or authors of the Document to the Document's overall - subject (or to related matters) and contains nothing that could - fall directly within that overall subject. (Thus, if the Document - is in part a textbook of mathematics, a Secondary Section may not - explain any mathematics.) The relationship could be a matter of - historical connection with the subject or with related matters, or - of legal, commercial, philosophical, ethical or political position - regarding them. - - The "Invariant Sections" are certain Secondary Sections whose - titles are designated, as being those of Invariant Sections, in the - notice that says that the Document is released under this License. - If a section does not fit the above definition of Secondary then it - is not allowed to be designated as Invariant. The Document may - contain zero Invariant Sections. If the Document does not identify - any Invariant Sections then there are none. - - The "Cover Texts" are certain short passages of text that are - listed, as Front-Cover Texts or Back-Cover Texts, in the notice - that says that the Document is released under this License. - A Front-Cover Text may be at most 5 words, and a Back-Cover Text - may be at most 25 words. - - A "Transparent" copy of the Document means a machine-readable copy, - represented in a format whose specification is available to the - general public, that is suitable for revising the document - straightforwardly with generic text editors or (for images composed - of pixels) generic paint programs or (for drawings) some widely - available drawing editor, and that is suitable for input to text - formatters or for automatic translation to a variety of formats - suitable for input to text formatters. A copy made in an otherwise - Transparent file format whose markup, or absence of markup, has - been arranged to thwart or discourage subsequent modification by - readers is not Transparent. An image format is not Transparent if - used for any substantial amount of text. A copy that is not - "Transparent" is called "Opaque". - - Examples of suitable formats for Transparent copies include plain - ASCII without markup, Texinfo input format, LaTeX input format, - SGML or XML using a publicly available DTD, and standard-conforming - simple HTML, PostScript or PDF designed for human modification. - Examples of transparent image formats include PNG, XCF and JPG. - Opaque formats include proprietary formats that can be read and - edited only by proprietary word processors, SGML or XML for which - the DTD and/or processing tools are not generally available, and - the machine-generated HTML, PostScript or PDF produced by some word - processors for output purposes only. - - The "Title Page" means, for a printed book, the title page itself, - plus such following pages as are needed to hold, legibly, the - material this License requires to appear in the title page. For - works in formats which do not have any title page as such, "Title - Page" means the text near the most prominent appearance of the - work's title, preceding the beginning of the body of the text. - - The "publisher" means any person or entity that distributes copies - of the Document to the public. - - A section "Entitled XYZ" means a named subunit of the Document - whose title either is precisely XYZ or contains XYZ in parentheses - following text that translates XYZ in another language. (Here XYZ - stands for a specific section name mentioned below, such as - "Acknowledgements", "Dedications", "Endorsements", or "History".) - To "Preserve the Title" of such a section when you modify the - Document means that it remains a section "Entitled XYZ" according - to this definition. - - The Document may include Warranty Disclaimers next to the notice - which states that this License applies to the Document. These - Warranty Disclaimers are considered to be included by reference in - this License, but only as regards disclaiming warranties: any other - implication that these Warranty Disclaimers may have is void and - has no effect on the meaning of this License. - -2. VERBATIM COPYING - - You may copy and distribute the Document in any medium, either - commercially or noncommercially, provided that this License, the - copyright notices, and the license notice saying this License - applies to the Document are reproduced in all copies, and that you - add no other conditions whatsoever to those of this License. You - may not use technical measures to obstruct or control the reading - or further copying of the copies you make or distribute. However, - you may accept compensation in exchange for copies. If you - distribute a large enough number of copies you must also follow the - conditions in section 3. - - You may also lend copies, under the same conditions stated above, - and you may publicly display copies. - -3. COPYING IN QUANTITY - - If you publish printed copies (or copies in media that commonly - have printed covers) of the Document, numbering more than 100, and - the Document's license notice requires Cover Texts, you must - enclose the copies in covers that carry, clearly and legibly, all - these Cover Texts: Front-Cover Texts on the front cover, and - Back-Cover Texts on the back cover. Both covers must also clearly - and legibly identify you as the publisher of these copies. The - front cover must present the full title with all words of the title - equally prominent and visible. You may add other material on the - covers in addition. Copying with changes limited to the covers, as - long as they preserve the title of the Document and satisfy these - conditions, can be treated as verbatim copying in other respects. - - If the required texts for either cover are too voluminous to fit - legibly, you should put the first ones listed (as many as fit - reasonably) on the actual cover, and continue the rest onto - adjacent pages. - - If you publish or distribute Opaque copies of the Document - numbering more than 100, you must either include a machine-readable - Transparent copy along with each Opaque copy, or state in or with - each Opaque copy a computer-network location from which the general - network-using public has access to download using public-standard - network protocols a complete Transparent copy of the Document, free - of added material. If you use the latter option, you must take - reasonably prudent steps, when you begin distribution of Opaque - copies in quantity, to ensure that this Transparent copy will - remain thus accessible at the stated location until at least one - year after the last time you distribute an Opaque copy (directly or - through your agents or retailers) of that edition to the public. - - It is requested, but not required, that you contact the authors of - the Document well before redistributing any large number of copies, - to give them a chance to provide you with an updated version of the - Document. - -4. MODIFICATIONS - - You may copy and distribute a Modified Version of the Document - under the conditions of sections 2 and 3 above, provided that you - release the Modified Version under precisely this License, with the - Modified Version filling the role of the Document, thus licensing - distribution and modification of the Modified Version to whoever - possesses a copy of it. In addition, you must do these things in - the Modified Version: - - #+attr_texinfo: :enum A - 1. Use in the Title Page (and on the covers, if any) a title - distinct from that of the Document, and from those of previous - versions (which should, if there were any, be listed in the - History section of the Document). You may use the same title as - a previous version if the original publisher of that version - gives permission. - - 2. List on the Title Page, as authors, one or more persons or - entities responsible for authorship of the modifications in the - Modified Version, together with at least five of the principal - authors of the Document (all of its principal authors, if it has - fewer than five), unless they release you from this requirement. - - 3. State on the Title page the name of the publisher of the - Modified Version, as the publisher. - - 4. Preserve all the copyright notices of the Document. - - 5. Add an appropriate copyright notice for your modifications - adjacent to the other copyright notices. - - 6. Include, immediately after the copyright notices, a license - notice giving the public permission to use the Modified Version - under the terms of this License, in the form shown in the - Addendum below. - - 7. Preserve in that license notice the full lists of Invariant - Sections and required Cover Texts given in the Document's - license notice. - - 8. Include an unaltered copy of this License. - - 9. Preserve the section Entitled "History", Preserve its Title, and - add to it an item stating at least the title, year, new authors, - and publisher of the Modified Version as given on the Title - Page. If there is no section Entitled "History" in the Document, - create one stating the title, year, authors, and publisher of - the Document as given on its Title Page, then add an item - describing the Modified Version as stated in the previous - sentence. - - 10. Preserve the network location, if any, given in the Document - for public access to a Transparent copy of the Document, and - likewise the network locations given in the Document for - previous versions it was based on. These may be placed in the - "History" section. You may omit a network location for a work - that was published at least four years before the Document - itself, or if the original publisher of the version it refers - to gives permission. - - 11. For any section Entitled "Acknowledgements" or "Dedications", - Preserve the Title of the section, and preserve in the section - all the substance and tone of each of the contributor - acknowledgements and/or dedications given therein. - - 12. Preserve all the Invariant Sections of the Document, unaltered - in their text and in their titles. Section numbers or the - equivalent are not considered part of the section titles. - - 13. Delete any section Entitled "Endorsements". Such a section may - not be included in the Modified Version. - - 14. Do not retitle any existing section to be Entitled - "Endorsements" or to conflict in title with any Invariant - Section. - - 15. Preserve any Warranty Disclaimers. - - If the Modified Version includes new front-matter sections or - appendices that qualify as Secondary Sections and contain no material - copied from the Document, you may at your option designate some or all - of these sections as invariant. To do this, add their titles to the - list of Invariant Sections in the Modified Version's license notice. - These titles must be distinct from any other section titles. - - You may add a section Entitled "Endorsements", provided it contains - nothing but endorsements of your Modified Version by various - parties---for example, statements of peer review or that the text has - been approved by an organization as the authoritative definition of a - standard. - - You may add a passage of up to five words as a Front-Cover Text, and a - passage of up to 25 words as a Back-Cover Text, to the end of the list - of Cover Texts in the Modified Version. Only one passage of - Front-Cover Text and one of Back-Cover Text may be added by (or - through arrangements made by) any one entity. If the Document already - includes a cover text for the same cover, previously added by you or - by arrangement made by the same entity you are acting on behalf of, - you may not add another; but you may replace the old one, on explicit - permission from the previous publisher that added the old one. - - The author(s) and publisher(s) of the Document do not by this License - give permission to use their names for publicity for or to assert or - imply endorsement of any Modified Version. - -5. COMBINING DOCUMENTS - - You may combine the Document with other documents released under - this License, under the terms defined in section 4 above for - modified versions, provided that you include in the combination all - of the Invariant Sections of all of the original documents, - unmodified, and list them all as Invariant Sections of your - combined work in its license notice, and that you preserve all - their Warranty Disclaimers. - - The combined work need only contain one copy of this License, and - multiple identical Invariant Sections may be replaced with a single - copy. If there are multiple Invariant Sections with the same name - but different contents, make the title of each such section unique - by adding at the end of it, in parentheses, the name of the - original author or publisher of that section if known, or else - a unique number. Make the same adjustment to the section titles in - the list of Invariant Sections in the license notice of the - combined work. - - In the combination, you must combine any sections Entitled - "History" in the various original documents, forming one section - Entitled "History"; likewise combine any sections Entitled - "Acknowledgements", and any sections Entitled "Dedications". You - must delete all sections Entitled "Endorsements." - -6. COLLECTIONS OF DOCUMENTS - - You may make a collection consisting of the Document and other - documents released under this License, and replace the individual - copies of this License in the various documents with a single copy - that is included in the collection, provided that you follow the - rules of this License for verbatim copying of each of the documents - in all other respects. - - You may extract a single document from such a collection, and - distribute it individually under this License, provided you insert - a copy of this License into the extracted document, and follow this - License in all other respects regarding verbatim copying of that - document. - -7. AGGREGATION WITH INDEPENDENT WORKS - - A compilation of the Document or its derivatives with other - separate and independent documents or works, in or on a volume of - a storage or distribution medium, is called an "aggregate" if the - copyright resulting from the compilation is not used to limit the - legal rights of the compilation's users beyond what the individual - works permit. When the Document is included in an aggregate, this - License does not apply to the other works in the aggregate which - are not themselves derivative works of the Document. - - If the Cover Text requirement of section 3 is applicable to these - copies of the Document, then if the Document is less than one half - of the entire aggregate, the Document's Cover Texts may be placed - on covers that bracket the Document within the aggregate, or the - electronic equivalent of covers if the Document is in electronic - form. Otherwise they must appear on printed covers that bracket - the whole aggregate. - -8. TRANSLATION - - Translation is considered a kind of modification, so you may - distribute translations of the Document under the terms of - section 4. Replacing Invariant Sections with translations requires - special permission from their copyright holders, but you may - include translations of some or all Invariant Sections in addition - to the original versions of these Invariant Sections. You may - include a translation of this License, and all the license notices - in the Document, and any Warranty Disclaimers, provided that you - also include the original English version of this License and the - original versions of those notices and disclaimers. In case of - a disagreement between the translation and the original version of - this License or a notice or disclaimer, the original version will - prevail. - - If a section in the Document is Entitled "Acknowledgements", - "Dedications", or "History", the requirement (section 4) to - Preserve its Title (section 1) will typically require changing the - actual title. - -9. TERMINATION - - You may not copy, modify, sublicense, or distribute the Document - except as expressly provided under this License. Any attempt - otherwise to copy, modify, sublicense, or distribute it is void, - and will automatically terminate your rights under this License. - - However, if you cease all violation of this License, then your - license from a particular copyright holder is reinstated (a) - provisionally, unless and until the copyright holder explicitly and - finally terminates your license, and (b) permanently, if the - copyright holder fails to notify you of the violation by some - reasonable means prior to 60 days after the cessation. - - Moreover, your license from a particular copyright holder is - reinstated permanently if the copyright holder notifies you of the - violation by some reasonable means, this is the first time you have - received notice of violation of this License (for any work) from - that copyright holder, and you cure the violation prior to 30 days - after your receipt of the notice. - - Termination of your rights under this section does not terminate - the licenses of parties who have received copies or rights from you - under this License. If your rights have been terminated and not - permanently reinstated, receipt of a copy of some or all of the - same material does not give you any rights to use it. - -10. FUTURE REVISIONS OF THIS LICENSE - - The Free Software Foundation may publish new, revised versions of - the GNU Free Documentation License from time to time. Such new - versions will be similar in spirit to the present version, but may - differ in detail to address new problems or concerns. See - https://www.gnu.org/copyleft/. - - Each version of the License is given a distinguishing version - number. If the Document specifies that a particular numbered - version of this License "or any later version" applies to it, you - have the option of following the terms and conditions either of - that specified version or of any later version that has been - published (not as a draft) by the Free Software Foundation. If - the Document does not specify a version number of this License, - you may choose any version ever published (not as a draft) by the - Free Software Foundation. If the Document specifies that a proxy - can decide which future versions of this License can be used, that - proxy's public statement of acceptance of a version permanently - authorizes you to choose that version for the Document. - -11. RELICENSING - - "Massive Multiauthor Collaboration Site" (or "MMC Site") means any - World Wide Web server that publishes copyrightable works and also - provides prominent facilities for anybody to edit those works. - A public wiki that anybody can edit is an example of such - a server. A "Massive Multiauthor Collaboration" (or "MMC") - contained in the site means any set of copyrightable works thus - published on the MMC site. - - "CC-BY-SA" means the Creative Commons Attribution-Share Alike 3.0 - license published by Creative Commons Corporation, - a not-for-profit corporation with a principal place of business in - San Francisco, California, as well as future copyleft versions of - that license published by that same organization. - - "Incorporate" means to publish or republish a Document, in whole - or in part, as part of another Document. - - An MMC is "eligible for relicensing" if it is licensed under this - License, and if all works that were first published under this - License somewhere other than this MMC, and subsequently - incorporated in whole or in part into the MMC, (1) had no cover - texts or invariant sections, and (2) were thus incorporated prior - to November 1, 2008. - - The operator of an MMC Site may republish an MMC contained in the - site under CC-BY-SA on the same site at any time before August 1, - 2009, provided the MMC is eligible for relicensing. - -#+texinfo: @page - -* ADDENDUM: How to use this License for your documents -:PROPERTIES: -:UNNUMBERED: notoc -:END: - -To use this License in a document you have written, include a copy of -the License in the document and put the following copyright and -license notices just after the title page: - -#+begin_example - Copyright (C) YEAR YOUR NAME. - Permission is granted to copy, distribute and/or modify this document - under the terms of the GNU Free Documentation License, Version 1.3 - or any later version published by the Free Software Foundation; - with no Invariant Sections, no Front-Cover Texts, and no Back-Cover - Texts. A copy of the license is included in the section entitled ``GNU - Free Documentation License''. -#+end_example - -If you have Invariant Sections, Front-Cover Texts and Back-Cover Texts, -replace the "with...Texts."\nbsp{}line with this: - -#+begin_example - with the Invariant Sections being LIST THEIR TITLES, with - the Front-Cover Texts being LIST, and with the Back-Cover Texts - being LIST. -#+end_example - -If you have Invariant Sections without Cover Texts, or some other -combination of the three, merge those two alternatives to suit the -situation. - -If your document contains nontrivial examples of program code, we -recommend releasing these examples in parallel under your choice of -free software license, such as the GNU General Public License, to -permit their use in free software. diff --git a/docs/man.org b/docs/man.org index 279165c..36dbc42 100644 --- a/docs/man.org +++ b/docs/man.org @@ -12,7 +12,59 @@ vericert - A formally verified high-level synthesis tool. * DESCRIPTION -- -drtl :: Print intermediate RTL. +** HLS Options + +- --no-hls :: Do not use HLS and generate standard flow +- --simulate :: Simulate the result with the Verilog semantics +- --debug-hls :: Add debug logic to the Verilog +- --initialise-stack :: initialise the stack to all 0s + +** HLS Optimisations + +- -fschedule :: Schedule the resulting hardware [off] +- -fif-conversion :: If-conversion optimisation [off] + +** General options + +- -stdlib :: Set the path of the Compcert run-time library +- -v :: Print external commands before invoking them +- -timings :: Show the time spent in various compiler passes +- -version :: Print the version string and exit +- -target :: Generate code for the given target +- -conf :: Read configuration from file +- @ :: Read command line options from + +** Tracing Options + +- -dprepro :: Save C file after preprocessing in .i +- -dparse :: Save C file after parsing and elaboration in .parsed.c +- -dc :: Save generated C in .compcert.c +- -dclight :: Save generated Clight in .light.c +- -dcminor :: Save generated Cminor in .cm +- -drtl :: Save RTL at various optimization points in .rtl. +- -drtlblock :: Save RTLBlock .rtlblock +- -dhtl :: Save HTL before Verilog generation .htl +- -dltl :: Save LTL after register allocation in .ltl +- -dmach :: Save generated Mach code in .mach +- -dasm :: Save generated assembly in .s +- -dall :: Save all generated intermediate files in . +- -sdump :: Save info for post-linking validation in .json +- -o :: Generate output in + +** Diagnostic options + +- -Wall :: Enable all warnings +- -W :: Enable the specific +- -Wno- :: Disable the specific +- -Werror :: Make all warnings into errors +- -Werror= :: Turn into an error +- -Wno-error= :: Turn into a warning even if -Werror is specified +- -Wfatal-errors :: Turn all errors into fatal errors aborting the compilation +- -fdiagnostics-color :: Turn on colored diagnostics +- -fno-diagnostics-color :: Turn of colored diagnostics +- -fmax-errors= :: Maximum number of errors to report +- -fdiagnostics-show-option :: Print the option name with mappable diagnostics +- -fno-diagnostics-show-option :: Turn of printing of options with mappable diagnostics * AUTHOR diff --git a/docs/res/coqdoc.css b/docs/res/coqdoc.css new file mode 100644 index 0000000..5572aaa --- /dev/null +++ b/docs/res/coqdoc.css @@ -0,0 +1,867 @@ +body { + padding: 0px 0px; + margin: 0px 0px; + padding-left: 1em; + background-color: white; + font-family: sans-serif; + background-repeat: no-repeat; + background-size: 100%; +} + +#page { + display: block; + padding: 0px; + margin: 0px; +} + +#header { + min-height: 100px; + max-width: 760px; + margin: 0 auto; + padding-left: 80px; + padding-right: 80px; + padding-top: 30px; +} + +#header h1 { + padding: 0; + margin: 0; +} + +/* Menu */ +ul#menu { + padding: 0; + display: block; + margin: auto; +} + +ul#menu li, div.button { + display: inline-block; + background-color: white; + padding: 5px; + font-size: .70em; + text-transform: uppercase; + width: 30%; + text-align: center; + font-weight: 600; +} + +div.button { + margin-top:5px; + width: 40%; +} + +#button_block {margin-top:50px;} + +ul#menu a.hover li { + background-color: red; +} + +/* Contents */ + +#main, #main_home { + display: block; + padding: 80px; + padding-top: 60px; /* BCP */ + font-size: 100%; + line-height: 100%; + max-width: 760px; + background-color: #ffffff; + margin: 0 auto; +} + +#main_home { + background-color: rgba(255, 255, 255, 0.5); +} + +#index_content div.intro p { + font-size: 12px; +} + +#main h1 { + /* line-height: 80%; */ + line-height: normal; + padding-top: 3px; + padding-bottom: 4px; + /* Demitri: font-size: 22pt; */ + font-size: 200%; /* BCP */ +} + +/* allow for multi-line headers */ +#main a.idref:visited {color : #416DFF; text-decoration : none; } +#main a.idref:link {color : #416DFF; text-decoration : none; } +#main a.idref:hover {text-decoration : none; } +#main a.idref:active {text-decoration : none; } + +#main a.modref:visited {color : #416DFF; text-decoration : none; } +#main a.modref:link {color : #416DFF; text-decoration : none; } +#main a.modref:hover {text-decoration : none; } +#main a.modref:active {text-decoration : none; } + +#main .keyword { color : #697f2f } +#main { color: black } + +/* General section class - applies to all .section IDs */ +.section { + padding-top: 12px; + padding-bottom: 11px; + padding-left: 8px; + margin-top: 5px; + margin-bottom: 3px; + margin-top: 18px; + font-size : 125%; + color: #ffffff; +} + +/* Book title in header */ +.booktitleinheader { + color: #000000; + text-transform: uppercase; + font-weight: 300; + letter-spacing: 1px; + font-size: 125%; + margin-left: 0px; + margin-bottom: 22px; + } + +/* Chapter titles */ +.libtitle { + max-width: 860px; + text-transform: uppercase; + margin: 5px auto; + font-weight: 500; + padding-bottom: 2px; + font-size: 120%; + letter-spacing: 3px; + } + +.subtitle { + display: block; + padding-top: 10px; + font-size: 70%; + line-height: normal; +} + +h2.section { + color: #2a2c71; + background-color: transparent; + padding-left: 0px; + padding-top: 0px; + padding-bottom: 0px; + /* margin-top: 0px; */ + margin-top: 9px; /* BCP 2/20 */ + font-size : 135%; } + +h3.section { + /* background-color: rgb(90%,90%,100%); */ + background-color: white; + /* padding-left: 8px; */ + padding-left: 0px; + /* padding-top: 7px; */ + padding-top: 0px; + /* padding-bottom: 0px; */ + padding-bottom: 0.5em; + /* margin-top: 9px; */ + /* margin-top: 4px; (BCP 2/20) */ + margin-top: 9px; /* BCP 2/20 */ + font-size : 115%; + color:black; +} + +h4.section { + margin-top: .2em; + background-color: white; + color: #2a2c71; + padding-left: 0px; + padding-top: 0px; + padding-bottom: 0.5em; /* 0px; */ + font-size : 100%; + font-style : bold; + text-decoration : underline; +} + +#header p { + font-size: 13px; +} + +/* Sets up Main ID and margins */ + +#main .doc { + margin: 0px auto; + font-size: 14px; + line-height: 22px; + /* max-width: 570px; */ + color: black; + /* text-align: justify; */ + border-style: plain; + /* This might work better after changing back to standard coqdoc: */ + padding-top: 10px; + /* padding-top: 18px; */ +} + +.quote { + margin-left: auto; + margin-right: auto; + width: 40em; + color: darkred; +} + +.loud { + color: darkred; +} + +pre { + margin-top: 0px; + margin-bottom: 0px; +} + +.inlinecode { + display: inline; + /* font-size: 125%; */ + color: #444444; + font-family: monospace } + +.doc .inlinecode { + display: inline; + font-size: 105%; + color: rgb(35%,35%,70%); + font-family: monospace } + +.doc .inlinecode .id { +/* I am changing this to white in style below: + color: rgb(35%,35%,70%); +*/ +} + +h1 .inlinecode .id, h1.section span.inlinecode { + color: #ffffff; +} + +.inlinecodenm { + display: inline; + /* font-size: 125%; */ + color: #444444; +} + +.doc .inlinecodenm { + display: inline; + color: rgb(35%,35%,70%); +} + +.doc .inlinecodenm .id { + color: rgb(35%,35%,70%); +} + +.doc .code { + display: inline; + font-size: 110%; + color: rgb(35%,35%,70%); + font-family: monospace; + padding-left: 0px; +} + +.comment { + display: inline; + font-family: monospace; + color: rgb(50%,50%,80%); +} + +.inlineref { + display: inline; + /* font-family: monospace; */ + color: rgb(50%,50%,80%); +} + +.show::before { + /* content: "more"; */ + content: "+"; +} + +.show { + background-color: rgb(93%,93%,93%); + display: inline; + font-family: monospace; + font-size: 60%; + padding-top: 1px; + padding-bottom: 2px; + padding-left: 4px; + padding-right: 4px; + color: rgb(60%,60%,60%); +} + +/* +.show { + display: inline; + font-family: monospace; + font-size: 60%; + padding-top: 0px; + padding-bottom: 0px; + padding-left: 10px; + border: 1px; + border-style: solid; + color: rgb(75%,75%,85%); +} +*/ + +.proofbox { + font-size: 90%; + color: rgb(75%,75%,75%); +} + +#main .less-space { + margin-top: -12px; +} + +/* Inline quizzes */ +.quiz:before { + color: rgb(40%,40%,40%); + /* content: "- Quick Check -" ; */ + display: block; + text-align: center; + margin-bottom: 5px; +} +.quiz { + border: 4px; + border-color: rgb(80%,80%,80%); + /* + margin-left: 40px; + margin-right: 100px; + */ + padding: 5px; + padding-left: 8px; + padding-right: 8px; + margin-top: 10px; + margin-bottom: 15px; + border-style: solid; +} + +/* For textual ones... */ +.show-old { + display: inline; + font-family: monospace; + font-size: 80%; + padding-top: 0px; + padding-bottom: 0px; + padding-left: 3px; + padding-right: 3px; + border: 1px; + margin-top: 5px; /* doesn't work?! */ + border-style: solid; + color: rgb(75%,75%,85%); +} + +.largebr { + margin-top: 10px; +} + +.code { + padding-left: 20px; + font-size: 14px; + font-family: monospace; + line-height: 17px; + margin-top: 9px; +} + +/* Working: +.code { + display: block; + padding-left: 0px; + font-size: 110%; + font-family: monospace; + } +*/ + +.code-space { + max-width: 50em; + margin-top: 0em; + /* margin-bottom: -.5em; */ + margin-left: auto; + margin-right: auto; +} + +.code-tight { + max-width: 50em; + margin-top: .6em; + /* margin-bottom: -.7em; */ + margin-left: auto; + margin-right: auto; +} + +hr.doublespaceincode { + height: 1pt; + visibility: hidden; + font-size: 10px; +} + +/* +code.br { + height: 5em; +} +*/ + +#main .citation { + color: rgb(70%,0%,0%); + text-decoration: underline; +} + +table.infrule { + border: 0px; + margin-left: 50px; + margin-top: .5em; + margin-bottom: 1.2em; +} + +td.infrule { + font-family: monospace; + text-align: center; + /* color: rgb(35%,35%,70%); */ + padding: 0px; + line-height: 100%; +} + +tr.infrulemiddle hr { + margin: 1px 0 1px 0; +} + +.infrulenamecol { + color: rgb(60%,60%,60%); + font-size: 80%; + padding-left: 1em; + padding-bottom: 0.1em +} + +#footer { + font-size: 65%; + font-family: sans-serif; +} + +.id { display: inline; } + +.id[title="constructor"] { + color: #697f2f; +} + +.id[title="var"], +.id[title="variable"] { + color: rgb(40%,0%,40%); +} + +.id[title="definition"] { + color: rgb(0%,40%,0%); +} + +.id[title="abbreviation"] { + color: rgb(0%,40%,0%); +} + +.id[title="lemma"] { + color: rgb(0%,40%,0%); +} + +.id[title="instance"] { + color: rgb(0%,40%,0%); +} + +.id[title="projection"] { + color: rgb(0%,40%,0%); +} + +.id[title="method"] { + color: rgb(0%,40%,0%); +} + +.id[title="inductive"] { + color: #034764; +} + +.id[title="record"] { + color: rgb(0%,0%,80%); +} + +.id[title="class"] { + color: rgb(0%,0%,80%); +} + +.id[title="keyword"] { + color : #697f2f; + /* color: black; */ +} + +.inlinecode .id { + color: rgb(0%,0%,0%); +} + +.nowrap { + white-space: nowrap; +} + +.HIDEFROMHTML { + display: none; +} + +/* TOC */ + +#toc h2 { +/* padding-top: 13px; */ + padding-bottom: 13px; + padding-left: 8px; + margin-top: 5px; + margin-top: 20px; + /* OLD: padding: 10px; + line-height: 120%; + background-color: rgb(60%,60%,100%); */ +} + +#toc h2.ui-accordion-header { + padding: .5em .5em .5em .7em; + outline: none; +} + +#toc .ui-accordion .ui-accordion-content { + padding: 0.5em 2.5em 0.8em .9em; + border-top: 0; + margin-bottom: 1em; + /* bottom rule */ + border: none; + border-bottom: 1px solid transparent; + transition: border-bottom-color 0.25s ease-in; + transition-delay: 0.15s; +} + +#toc .ui-accordion .ui-accordion-content-active { + border-bottom: 1px solid #9b9b9b; + transition-delay: 0s; +} + +#toc h2.ui-accordion-header-active { + background: silver !important; +} + +#toc h2:not(.ui-accordion-header-active):hover { + background: rgba(0,0,0,0.04) !important; +} + +#toc h2 a:hover { + text-decoration: underline; +} + +#toc h2:hover::after { + content: "expand ▾"; + font-size: 80%; + float: right; + margin-top: 0.2em; + color: silver; + opacity: 1; + transition: opacity .5s ease-in-out; +} + +#toc h2.ui-accordion-header-active:hover::after { + opacity: 0; +} + +#toc h2 .select { background-image: url('media/image/arrow_down.jpg'); } +div#sec1.hide { display: none; } + +#toc ul { + padding-top: 8px; + font-size: 14px; + padding-left: 0; +} + +#toc ul ul { + margin-bottom: -8px; +} + +#toc li { + font-weight: 600; + list-style-type: none; + padding-top: 12px; + padding-bottom: 8px; +} + +#toc li li { + font-weight: 300; + list-style-type: circle; + padding-bottom: 3px; + padding-top: 0px; + margin-left: 19px; +} + + + + +/* Accordion Overrides */ + +/* Widget Bar */ +.ui-state-default, +.ui-widget-content .ui-state-default, +.ui-widget-header .ui-state-default, +.ui-button, +/* We use html here because we need a greater specificity to make sure disabled + works properly when clicked or hovered */ +html .ui-button.ui-state-disabled:hover, +html .ui-button.ui-state-disabled:active { + border: none!important; + /* BCP 3/17: I like it better without the rules... + border-bottom: 1px solid silver!important; */ + background: white !important; + font-weight: normal; + color: #454545!important; + font-weight: 400!important; + margin-top: 0px!important; + +} + +/* Misc visuals +----------------------------------*/ + +/* Corner radius */ +.ui-corner-all, +.ui-corner-top, +.ui-corner-left, +.ui-corner-tl { + border-top-left-radius: 0px!important; +} + +.ui-corner-all, +.ui-corner-top, +.ui-corner-right, +.ui-corner-tr { + border-top-right-radius: 0px!important; +} + +.ui-corner-all, +.ui-corner-bottom, +.ui-corner-left, +.ui-corner-bl { + border-bottom-left-radius: 0px!important; +} + +.ui-corner-all, +.ui-corner-bottom, +.ui-corner-right, +.ui-corner-br { + border-bottom-right-radius: 0px!important; +} + +html .ui-button.ui-state-disabled:focus { + color: red!important; +} + +/* Remove Icon */ +.ui-icon { display: none!important; } + +/* Widget */ +.ui-widget-content { + border: 1px solid #9e9e9e; + border-bottom-color: #b2b2b2; +} + +.ui-widget-content { + background: #ffffff; + color: #333333; +} + + +/* Index */ + +#index { + margin: 0; + padding: 0; + width: 100%; + font-style : normal; +} + +#index #frontispiece { + margin: auto; + padding: 1em; + width: 700px; +} + +.booktitle { + font-size : 300%; line-height: 100%; font-style:bold; + color: white; + padding-top: 70px; + padding-bottom: 20px; } +.authors { font-size : 200%; + line-height: 115%; } +.moreauthors { font-size : 170% } +.buttons { font-size : 170%; + margin-left: auto; + margin-right: auto; + font-style : bold; + } + +/* Link colors never changes */ + +A:link, A:visited, A:active, A:hover { + text-decoration: none; + color: #555555 +} + +/* Special color for the chapter header */ + +.booktitleinheader A:visited, .booktitleinheader A:active, .booktitleinheader A:hover, .booktitleinheader A:link { + text-decoration: none; + color: black; +} + +#index #entrance { + text-align: center; +} + +/* This was removed via CSS but the HTML is still generated */ +#index #footer { + display: none; +} + +.paragraph { + height: 0.6em; +} + +ul.doclist { + margin-top: 0em; + margin-bottom: 0em; +} + +/* Index styles */ + +/* Styles the author box (Intro class) and With (column class) */ + +.column { + float:left; + width: 43%; + margin:0 10px; + text-align: left; + font-size: 15px; + line-height: 25px; + padding-right: 20px; + min-height: 340px; +} + +.smallauthors { + font-size: 19px; + line-height: 25px; +} + +.mediumauthors { + font-size: 23px; + line-height: 33px; +} + +.largeauthors { + font-size: 28px; + line-height: 40px; +} + +.intro { + width: 35%; + font-size: 21px; + line-height: 27px; + font-weight: 600; + padding-right: 20px; +} + +.column.pub { + width: 40%; + margin-bottom: 20px; +} + +#index_content { + width: 100%!important; + display: block; + min-height: 400px; +} + +div.column.pub table tbody tr td { + text-align: center; padding: 10px; +} +div.column.pub table tbody tr td p { + text-align: left; + margin-top: 0; + font-weight: 600; + font-size: 13px!important; + line-height: 18px; +} + +/* Tables */ + +td.tab { + height: 16px; + font-weight: 600; + padding-left: 5px; + text-align: left!important; +} + +/* Styles tables on the index -- body class sf_home is used there */ + +body.sf_home table { + min-height: 450px; + vertical-align: top; +} + +body.sf_home table td { + vertical-align: top; + +} +body.sf_home table td p { + min-height: 100px; + +} + +table.logical { background-color: rgba(144, 160, 209, 0.5); } +table.logical tbody tr td.tab { background-color: #91a1d1; } + +table.language_found { background-color: rgba(178, 88, 88, 0.5); } +table.language_found tbody tr td.tab { background-color: #b25959; } + +table.algo { background-color: rgba(194, 194, 108, 0.5); } +table.algo tbody tr td.tab { background-color: #c2c26c; } + +table.qc { background-color: rgba(185, 170, 185, 0.5); } +table.qc tbody tr td.tab { background-color: #8b7d95; } + +table.vc { background-color: rgba(159, 125, 140, 0.5); } +table.vc tbody tr td.tab { background-color: rgb(159, 125, 140); } + +table.slf { background-color: rgba(219, 178, 127, 0.5); } +table.slf tbody tr td.tab { background-color: rgb(219, 178, 127); } + +/* Suggested background color for next volume */ +/* #c07d62 */ + +.ui-draggable, .ui-droppable { + background-position: top; +} + +/* Chapter dependencies (SVG) */ +.deps a polygon:hover { opacity: 0.6; stroke-width: 2; } +.deps a text { pointer-events: none; } + +/* A few specific things for the top-level SF landing page */ + +body.sf_home {background-color: #ededed; background-image: url(../media/image/core_mem_bg.jpg); } + +body.sf_home #header { + background-image: url(../media/image/core_mem_hdr_bg.jpg); + padding-bottom: 20px; +} + +body.sf_home #main_home { + background-color: transparent; +} + +/* A partial fix to a coqdoc bug... + See https://github.com/DeepSpec/sfdev/issues/236 */ +.inlinecode { white-space: pre; } +.inlinecode br { display: none; } + +#header { background-color: rgba(190, 170, 190, 0.5); } + +/* This volume's color */ +.section, ul#menu li.section_name, div.button { background-color: #8b7d95; } + +.slide img { + border: 2px solid gray; + margin: 1em; +} diff --git a/docs/res/fdl.org b/docs/res/fdl.org new file mode 100644 index 0000000..81e2cd9 --- /dev/null +++ b/docs/res/fdl.org @@ -0,0 +1,489 @@ +# The GNU Free Documentation License. +#+begin_center +Version 1.3, 3 November 2008 +#+end_center + +# This file is intended to be included within another document. + +#+begin_verse +Copyright \copy{} 2000, 2001, 2002, 2007, 2008 Free Software Foundation, Inc. +https://fsf.org/ + +Everyone is permitted to copy and distribute verbatim copies +of this license document, but changing it is not allowed. +#+end_verse + +0. [@0] PREAMBLE + + The purpose of this License is to make a manual, textbook, or other + functional and useful document @@texinfo:@dfn{@@free@@texinfo:}@@ + in the sense of freedom: to assure everyone the effective freedom + to copy and redistribute it, with or without modifying it, either + commercially or noncommercially. Secondarily, this License + preserves for the author and publisher a way to get credit for + their work, while not being considered responsible for + modifications made by others. + + This License is a kind of "copyleft", which means that derivative + works of the document must themselves be free in the same sense. + It complements the GNU General Public License, which is a copyleft + license designed for free software. + + We have designed this License in order to use it for manuals for + free software, because free software needs free documentation: + a free program should come with manuals providing the same freedoms + that the software does. But this License is not limited to + software manuals; it can be used for any textual work, regardless + of subject matter or whether it is published as a printed book. We + recommend this License principally for works whose purpose is + instruction or reference. + +1. APPLICABILITY AND DEFINITIONS + + This License applies to any manual or other work, in any medium, + that contains a notice placed by the copyright holder saying it can + be distributed under the terms of this License. Such a notice + grants a world-wide, royalty-free license, unlimited in duration, + to use that work under the conditions stated herein. The + "Document", below, refers to any such manual or work. Any member + of the public is a licensee, and is addressed as "you". You accept + the license if you copy, modify or distribute the work in a way + requiring permission under copyright law. + + A "Modified Version" of the Document means any work containing the + Document or a portion of it, either copied verbatim, or with + modifications and/or translated into another language. + + A "Secondary Section" is a named appendix or a front-matter section + of the Document that deals exclusively with the relationship of the + publishers or authors of the Document to the Document's overall + subject (or to related matters) and contains nothing that could + fall directly within that overall subject. (Thus, if the Document + is in part a textbook of mathematics, a Secondary Section may not + explain any mathematics.) The relationship could be a matter of + historical connection with the subject or with related matters, or + of legal, commercial, philosophical, ethical or political position + regarding them. + + The "Invariant Sections" are certain Secondary Sections whose + titles are designated, as being those of Invariant Sections, in the + notice that says that the Document is released under this License. + If a section does not fit the above definition of Secondary then it + is not allowed to be designated as Invariant. The Document may + contain zero Invariant Sections. If the Document does not identify + any Invariant Sections then there are none. + + The "Cover Texts" are certain short passages of text that are + listed, as Front-Cover Texts or Back-Cover Texts, in the notice + that says that the Document is released under this License. + A Front-Cover Text may be at most 5 words, and a Back-Cover Text + may be at most 25 words. + + A "Transparent" copy of the Document means a machine-readable copy, + represented in a format whose specification is available to the + general public, that is suitable for revising the document + straightforwardly with generic text editors or (for images composed + of pixels) generic paint programs or (for drawings) some widely + available drawing editor, and that is suitable for input to text + formatters or for automatic translation to a variety of formats + suitable for input to text formatters. A copy made in an otherwise + Transparent file format whose markup, or absence of markup, has + been arranged to thwart or discourage subsequent modification by + readers is not Transparent. An image format is not Transparent if + used for any substantial amount of text. A copy that is not + "Transparent" is called "Opaque". + + Examples of suitable formats for Transparent copies include plain + ASCII without markup, Texinfo input format, LaTeX input format, + SGML or XML using a publicly available DTD, and standard-conforming + simple HTML, PostScript or PDF designed for human modification. + Examples of transparent image formats include PNG, XCF and JPG. + Opaque formats include proprietary formats that can be read and + edited only by proprietary word processors, SGML or XML for which + the DTD and/or processing tools are not generally available, and + the machine-generated HTML, PostScript or PDF produced by some word + processors for output purposes only. + + The "Title Page" means, for a printed book, the title page itself, + plus such following pages as are needed to hold, legibly, the + material this License requires to appear in the title page. For + works in formats which do not have any title page as such, "Title + Page" means the text near the most prominent appearance of the + work's title, preceding the beginning of the body of the text. + + The "publisher" means any person or entity that distributes copies + of the Document to the public. + + A section "Entitled XYZ" means a named subunit of the Document + whose title either is precisely XYZ or contains XYZ in parentheses + following text that translates XYZ in another language. (Here XYZ + stands for a specific section name mentioned below, such as + "Acknowledgements", "Dedications", "Endorsements", or "History".) + To "Preserve the Title" of such a section when you modify the + Document means that it remains a section "Entitled XYZ" according + to this definition. + + The Document may include Warranty Disclaimers next to the notice + which states that this License applies to the Document. These + Warranty Disclaimers are considered to be included by reference in + this License, but only as regards disclaiming warranties: any other + implication that these Warranty Disclaimers may have is void and + has no effect on the meaning of this License. + +2. VERBATIM COPYING + + You may copy and distribute the Document in any medium, either + commercially or noncommercially, provided that this License, the + copyright notices, and the license notice saying this License + applies to the Document are reproduced in all copies, and that you + add no other conditions whatsoever to those of this License. You + may not use technical measures to obstruct or control the reading + or further copying of the copies you make or distribute. However, + you may accept compensation in exchange for copies. If you + distribute a large enough number of copies you must also follow the + conditions in section 3. + + You may also lend copies, under the same conditions stated above, + and you may publicly display copies. + +3. COPYING IN QUANTITY + + If you publish printed copies (or copies in media that commonly + have printed covers) of the Document, numbering more than 100, and + the Document's license notice requires Cover Texts, you must + enclose the copies in covers that carry, clearly and legibly, all + these Cover Texts: Front-Cover Texts on the front cover, and + Back-Cover Texts on the back cover. Both covers must also clearly + and legibly identify you as the publisher of these copies. The + front cover must present the full title with all words of the title + equally prominent and visible. You may add other material on the + covers in addition. Copying with changes limited to the covers, as + long as they preserve the title of the Document and satisfy these + conditions, can be treated as verbatim copying in other respects. + + If the required texts for either cover are too voluminous to fit + legibly, you should put the first ones listed (as many as fit + reasonably) on the actual cover, and continue the rest onto + adjacent pages. + + If you publish or distribute Opaque copies of the Document + numbering more than 100, you must either include a machine-readable + Transparent copy along with each Opaque copy, or state in or with + each Opaque copy a computer-network location from which the general + network-using public has access to download using public-standard + network protocols a complete Transparent copy of the Document, free + of added material. If you use the latter option, you must take + reasonably prudent steps, when you begin distribution of Opaque + copies in quantity, to ensure that this Transparent copy will + remain thus accessible at the stated location until at least one + year after the last time you distribute an Opaque copy (directly or + through your agents or retailers) of that edition to the public. + + It is requested, but not required, that you contact the authors of + the Document well before redistributing any large number of copies, + to give them a chance to provide you with an updated version of the + Document. + +4. MODIFICATIONS + + You may copy and distribute a Modified Version of the Document + under the conditions of sections 2 and 3 above, provided that you + release the Modified Version under precisely this License, with the + Modified Version filling the role of the Document, thus licensing + distribution and modification of the Modified Version to whoever + possesses a copy of it. In addition, you must do these things in + the Modified Version: + + #+attr_texinfo: :enum A + 1. Use in the Title Page (and on the covers, if any) a title + distinct from that of the Document, and from those of previous + versions (which should, if there were any, be listed in the + History section of the Document). You may use the same title as + a previous version if the original publisher of that version + gives permission. + + 2. List on the Title Page, as authors, one or more persons or + entities responsible for authorship of the modifications in the + Modified Version, together with at least five of the principal + authors of the Document (all of its principal authors, if it has + fewer than five), unless they release you from this requirement. + + 3. State on the Title page the name of the publisher of the + Modified Version, as the publisher. + + 4. Preserve all the copyright notices of the Document. + + 5. Add an appropriate copyright notice for your modifications + adjacent to the other copyright notices. + + 6. Include, immediately after the copyright notices, a license + notice giving the public permission to use the Modified Version + under the terms of this License, in the form shown in the + Addendum below. + + 7. Preserve in that license notice the full lists of Invariant + Sections and required Cover Texts given in the Document's + license notice. + + 8. Include an unaltered copy of this License. + + 9. Preserve the section Entitled "History", Preserve its Title, and + add to it an item stating at least the title, year, new authors, + and publisher of the Modified Version as given on the Title + Page. If there is no section Entitled "History" in the Document, + create one stating the title, year, authors, and publisher of + the Document as given on its Title Page, then add an item + describing the Modified Version as stated in the previous + sentence. + + 10. Preserve the network location, if any, given in the Document + for public access to a Transparent copy of the Document, and + likewise the network locations given in the Document for + previous versions it was based on. These may be placed in the + "History" section. You may omit a network location for a work + that was published at least four years before the Document + itself, or if the original publisher of the version it refers + to gives permission. + + 11. For any section Entitled "Acknowledgements" or "Dedications", + Preserve the Title of the section, and preserve in the section + all the substance and tone of each of the contributor + acknowledgements and/or dedications given therein. + + 12. Preserve all the Invariant Sections of the Document, unaltered + in their text and in their titles. Section numbers or the + equivalent are not considered part of the section titles. + + 13. Delete any section Entitled "Endorsements". Such a section may + not be included in the Modified Version. + + 14. Do not retitle any existing section to be Entitled + "Endorsements" or to conflict in title with any Invariant + Section. + + 15. Preserve any Warranty Disclaimers. + + If the Modified Version includes new front-matter sections or + appendices that qualify as Secondary Sections and contain no material + copied from the Document, you may at your option designate some or all + of these sections as invariant. To do this, add their titles to the + list of Invariant Sections in the Modified Version's license notice. + These titles must be distinct from any other section titles. + + You may add a section Entitled "Endorsements", provided it contains + nothing but endorsements of your Modified Version by various + parties---for example, statements of peer review or that the text has + been approved by an organization as the authoritative definition of a + standard. + + You may add a passage of up to five words as a Front-Cover Text, and a + passage of up to 25 words as a Back-Cover Text, to the end of the list + of Cover Texts in the Modified Version. Only one passage of + Front-Cover Text and one of Back-Cover Text may be added by (or + through arrangements made by) any one entity. If the Document already + includes a cover text for the same cover, previously added by you or + by arrangement made by the same entity you are acting on behalf of, + you may not add another; but you may replace the old one, on explicit + permission from the previous publisher that added the old one. + + The author(s) and publisher(s) of the Document do not by this License + give permission to use their names for publicity for or to assert or + imply endorsement of any Modified Version. + +5. COMBINING DOCUMENTS + + You may combine the Document with other documents released under + this License, under the terms defined in section 4 above for + modified versions, provided that you include in the combination all + of the Invariant Sections of all of the original documents, + unmodified, and list them all as Invariant Sections of your + combined work in its license notice, and that you preserve all + their Warranty Disclaimers. + + The combined work need only contain one copy of this License, and + multiple identical Invariant Sections may be replaced with a single + copy. If there are multiple Invariant Sections with the same name + but different contents, make the title of each such section unique + by adding at the end of it, in parentheses, the name of the + original author or publisher of that section if known, or else + a unique number. Make the same adjustment to the section titles in + the list of Invariant Sections in the license notice of the + combined work. + + In the combination, you must combine any sections Entitled + "History" in the various original documents, forming one section + Entitled "History"; likewise combine any sections Entitled + "Acknowledgements", and any sections Entitled "Dedications". You + must delete all sections Entitled "Endorsements." + +6. COLLECTIONS OF DOCUMENTS + + You may make a collection consisting of the Document and other + documents released under this License, and replace the individual + copies of this License in the various documents with a single copy + that is included in the collection, provided that you follow the + rules of this License for verbatim copying of each of the documents + in all other respects. + + You may extract a single document from such a collection, and + distribute it individually under this License, provided you insert + a copy of this License into the extracted document, and follow this + License in all other respects regarding verbatim copying of that + document. + +7. AGGREGATION WITH INDEPENDENT WORKS + + A compilation of the Document or its derivatives with other + separate and independent documents or works, in or on a volume of + a storage or distribution medium, is called an "aggregate" if the + copyright resulting from the compilation is not used to limit the + legal rights of the compilation's users beyond what the individual + works permit. When the Document is included in an aggregate, this + License does not apply to the other works in the aggregate which + are not themselves derivative works of the Document. + + If the Cover Text requirement of section 3 is applicable to these + copies of the Document, then if the Document is less than one half + of the entire aggregate, the Document's Cover Texts may be placed + on covers that bracket the Document within the aggregate, or the + electronic equivalent of covers if the Document is in electronic + form. Otherwise they must appear on printed covers that bracket + the whole aggregate. + +8. TRANSLATION + + Translation is considered a kind of modification, so you may + distribute translations of the Document under the terms of + section 4. Replacing Invariant Sections with translations requires + special permission from their copyright holders, but you may + include translations of some or all Invariant Sections in addition + to the original versions of these Invariant Sections. You may + include a translation of this License, and all the license notices + in the Document, and any Warranty Disclaimers, provided that you + also include the original English version of this License and the + original versions of those notices and disclaimers. In case of + a disagreement between the translation and the original version of + this License or a notice or disclaimer, the original version will + prevail. + + If a section in the Document is Entitled "Acknowledgements", + "Dedications", or "History", the requirement (section 4) to + Preserve its Title (section 1) will typically require changing the + actual title. + +9. TERMINATION + + You may not copy, modify, sublicense, or distribute the Document + except as expressly provided under this License. Any attempt + otherwise to copy, modify, sublicense, or distribute it is void, + and will automatically terminate your rights under this License. + + However, if you cease all violation of this License, then your + license from a particular copyright holder is reinstated (a) + provisionally, unless and until the copyright holder explicitly and + finally terminates your license, and (b) permanently, if the + copyright holder fails to notify you of the violation by some + reasonable means prior to 60 days after the cessation. + + Moreover, your license from a particular copyright holder is + reinstated permanently if the copyright holder notifies you of the + violation by some reasonable means, this is the first time you have + received notice of violation of this License (for any work) from + that copyright holder, and you cure the violation prior to 30 days + after your receipt of the notice. + + Termination of your rights under this section does not terminate + the licenses of parties who have received copies or rights from you + under this License. If your rights have been terminated and not + permanently reinstated, receipt of a copy of some or all of the + same material does not give you any rights to use it. + +10. FUTURE REVISIONS OF THIS LICENSE + + The Free Software Foundation may publish new, revised versions of + the GNU Free Documentation License from time to time. Such new + versions will be similar in spirit to the present version, but may + differ in detail to address new problems or concerns. See + https://www.gnu.org/copyleft/. + + Each version of the License is given a distinguishing version + number. If the Document specifies that a particular numbered + version of this License "or any later version" applies to it, you + have the option of following the terms and conditions either of + that specified version or of any later version that has been + published (not as a draft) by the Free Software Foundation. If + the Document does not specify a version number of this License, + you may choose any version ever published (not as a draft) by the + Free Software Foundation. If the Document specifies that a proxy + can decide which future versions of this License can be used, that + proxy's public statement of acceptance of a version permanently + authorizes you to choose that version for the Document. + +11. RELICENSING + + "Massive Multiauthor Collaboration Site" (or "MMC Site") means any + World Wide Web server that publishes copyrightable works and also + provides prominent facilities for anybody to edit those works. + A public wiki that anybody can edit is an example of such + a server. A "Massive Multiauthor Collaboration" (or "MMC") + contained in the site means any set of copyrightable works thus + published on the MMC site. + + "CC-BY-SA" means the Creative Commons Attribution-Share Alike 3.0 + license published by Creative Commons Corporation, + a not-for-profit corporation with a principal place of business in + San Francisco, California, as well as future copyleft versions of + that license published by that same organization. + + "Incorporate" means to publish or republish a Document, in whole + or in part, as part of another Document. + + An MMC is "eligible for relicensing" if it is licensed under this + License, and if all works that were first published under this + License somewhere other than this MMC, and subsequently + incorporated in whole or in part into the MMC, (1) had no cover + texts or invariant sections, and (2) were thus incorporated prior + to November 1, 2008. + + The operator of an MMC Site may republish an MMC contained in the + site under CC-BY-SA on the same site at any time before August 1, + 2009, provided the MMC is eligible for relicensing. + +#+texinfo: @page + +* ADDENDUM: How to use this License for your documents +:PROPERTIES: +:UNNUMBERED: notoc +:END: + +To use this License in a document you have written, include a copy of +the License in the document and put the following copyright and +license notices just after the title page: + +#+begin_example + Copyright (C) YEAR YOUR NAME. + Permission is granted to copy, distribute and/or modify this document + under the terms of the GNU Free Documentation License, Version 1.3 + or any later version published by the Free Software Foundation; + with no Invariant Sections, no Front-Cover Texts, and no Back-Cover + Texts. A copy of the license is included in the section entitled ``GNU + Free Documentation License''. +#+end_example + +If you have Invariant Sections, Front-Cover Texts and Back-Cover Texts, +replace the "with...Texts."\nbsp{}line with this: + +#+begin_example + with the Invariant Sections being LIST THEIR TITLES, with + the Front-Cover Texts being LIST, and with the Back-Cover Texts + being LIST. +#+end_example + +If you have Invariant Sections without Cover Texts, or some other +combination of the three, merge those two alternatives to suit the +situation. + +If your document contains nontrivial examples of program code, we +recommend releasing these examples in parallel under your choice of +free software license, such as the GNU General Public License, to +permit their use in free software. -- 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/Compiler.v | 102 ++++++++++++++++++------------------------------- src/hls/IfConversion.v | 8 +--- src/hls/RTLPargen.v | 20 +++------- 3 files changed, 45 insertions(+), 85 deletions(-) diff --git a/src/Compiler.v b/src/Compiler.v index b907629..7aef71c 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -1,7 +1,3 @@ -(*| -.. coq:: none -|*) - (* * Vericert: Verified high-level synthesis. * Copyright (C) 2019-2020 Yann Herklotz @@ -20,21 +16,15 @@ * along with this program. If not, see . *) -(*| -============== -Compiler Proof -============== - -.. contents:: Table of Contents - :depth: 2 +(** * Compiler Proof -This is the top-level module of the correctness proof and proves the final backwards simulation correct. +This is the top-level module of the correctness proof and proves the final backwards simulation +correct. -Imports -======= +** Imports -We first need to import all of the modules that are used in the correctness proof, which includes all of the passes that are performed in Vericert and the CompCert front end. -|*) +We first need to import all of the modules that are used in the correctness proof, which includes +all of the passes that are performed in Vericert and the CompCert front end. *) Require compcert.backend.Selection. Require compcert.backend.RTL. @@ -75,12 +65,10 @@ Require vericert.hls.Memorygen. Require Import vericert.hls.HTLgenproof. -(*| -Declarations -============ +(** ** Declarations -We then need to declare the external OCaml functions used to print out intermediate steps in the compilation, such as ``print_RTL``, ``print_HTL`` and ``print_RTLBlock``. -|*) +We then need to declare the external OCaml functions used to print out intermediate steps in the +compilation, such as ``print_RTL``, ``print_HTL`` and ``print_RTLBlock``. *) Parameter print_RTL: Z -> RTL.program -> unit. Parameter print_HTL: Z -> HTL.program -> unit. @@ -98,18 +86,16 @@ Proof. intros; unfold print. destruct (printer prog); auto. Qed. -(*| -We also declare some new notation, which is also used in CompCert to combine the monadic results of each pass. -|*) +(** We also declare some new notation, which is also used in CompCert to combine the monadic results +of each pass. *) Notation "a @@@ b" := (Compiler.apply_partial _ _ a b) (at level 50, left associativity). Notation "a @@ b" := (Compiler.apply_total _ _ a b) (at level 50, left associativity). -(*| -As printing is used in the translation but does not change the output, we need to prove that it has no effect so that it can be removed during the proof. -|*) +(** As printing is used in the translation but does not change the output, we need to prove that it +has no effect so that it can be removed during the proof. *) Lemma compose_print_identity: forall (A: Type) (x: res A) (f: A -> unit), @@ -118,9 +104,8 @@ Proof. intros. destruct x; simpl. rewrite print_identity. auto. auto. Qed. -(*| -Finally, some optimisation passes are only activated by a flag, which is handled by the following functions for partial and total passes. -|*) +(** Finally, some optimisation passes are only activated by a flag, which is handled by the following +functions for partial and total passes. *) Definition total_if {A: Type} (flag: unit -> bool) (f: A -> A) (prog: A) : A := @@ -171,12 +156,11 @@ Proof. intros. unfold match_if in *. destruct (flag tt). eauto. subst. apply forward_simulation_identity. Qed. -(*| -Top-level Translation ---------------------- +(** *** Top-level Translation -An optimised transformation function from ``RTL`` to ``Verilog`` can then be defined, which applies the front end compiler optimisations of CompCert to the RTL that is generated and then performs the two Vericert passes from RTL to HTL and then from HTL to Verilog. -|*) +An optimised transformation function from ``RTL`` to ``Verilog`` can then be defined, which applies +the front end compiler optimisations of CompCert to the RTL that is generated and then performs the +two Vericert passes from RTL to HTL and then from HTL to Verilog. *) Definition transf_backend (r : RTL.program) : res Verilog.program := OK r @@ -200,9 +184,8 @@ Definition transf_backend (r : RTL.program) : res Verilog.program := @@ print (print_HTL 1)*) @@ Veriloggen.transl_program. -(*| -The transformation functions from RTL to Verilog are then added to the backend of the CompCert transformations from Clight to RTL. -|*) +(** The transformation functions from RTL to Verilog are then added to the backend of the CompCert +transformations from Clight to RTL. *) Definition transf_hls (p : Csyntax.program) : res Verilog.program := OK p @@ -215,10 +198,6 @@ Definition transf_hls (p : Csyntax.program) : res Verilog.program := @@ print (print_RTL 0) @@@ transf_backend. -(*| -.. coq:: none -|*) - (* This is an unverified version of transf_hls with some experimental additions such as scheduling that aren't completed yet. *) Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := @@ -255,12 +234,10 @@ Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := @@ print (print_HTL 0) @@ Veriloggen.transl_program. -(*| -Correctness Proof -================= +(** ** Correctness Proof -Finally, the top-level definition for all the passes is defined, which combines the ``match_prog`` predicates of each translation pass from C until Verilog. -|*) +Finally, the top-level definition for all the passes is defined, which combines the ``match_prog`` +predicates of each translation pass from C until Verilog. *) Local Open Scope linking_scope. @@ -283,16 +260,14 @@ Definition CompCert's_passes := ::: mkpass Veriloggenproof.match_prog ::: pass_nil _. -(*| -These passes are then composed into a larger, top-level ``match_prog`` predicate which matches a C program directly with a Verilog program. -|*) +(** These passes are then composed into a larger, top-level ``match_prog`` predicate which matches a +C program directly with a Verilog program. *) Definition match_prog: Csyntax.program -> Verilog.program -> Prop := pass_match (compose_passes CompCert's_passes). -(*| -We then need to prove that this predicate holds, assuming that the translation is performed using the ``transf_hls`` function declared above. -|*) +(** We then need to prove that this predicate holds, assuming that the translation is performed +using the ``transf_hls`` function declared above. *) Theorem transf_hls_match: forall p tp, @@ -394,12 +369,12 @@ Ltac DestructM := apply Verilog.semantics_determinate. Qed. -(*| -Backward Simulation -------------------- +(** *** Backward Simulation -The following theorem is a *backward simulation* between the C and Verilog, which proves the semantics preservation of the translation. We can assume that the C and Verilog programs match, as we have proven previously in ``transf_hls_match`` that our translation implies that the ``match_prog`` predicate will hold. -|*) +The following theorem is a *backward simulation* between the C and Verilog, which proves the +semantics preservation of the translation. We can assume that the C and Verilog programs match, as +we have proven previously in ``transf_hls_match`` that our translation implies that the +``match_prog`` predicate will hold. *) Theorem c_semantic_preservation: forall p tp, @@ -416,9 +391,8 @@ Proof. exact (proj2 (cstrategy_semantic_preservation _ _ H)). Qed. -(*| -We can then use ``transf_hls_match`` to prove the backward simulation where the assumption is that the translation is performed using the ``transf_hls`` function and that it succeeds. -|*) +(** We can then use ``transf_hls_match`` to prove the backward simulation where the assumption is +that the translation is performed using the ``transf_hls`` function and that it succeeds. *) Theorem transf_c_program_correct: forall p tp, @@ -428,9 +402,9 @@ Proof. intros. apply c_semantic_preservation. apply transf_hls_match; auto. Qed. -(*| -The final theorem of the semantic preservation of the translation of separate translation units can also be proven correct, however, this is only because the translation fails if more than one translation unit is passed to Vericert at the moment. -|*) +(** The final theorem of the semantic preservation of the translation of separate translation units +can also be proven correct, however, this is only because the translation fails if more than one +translation unit is passed to Vericert at the moment. *) Theorem separate_transf_c_program_correct: forall c_units verilog_units c_program, 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(-) 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/Compiler.v | 18 +++++++++--------- src/hls/HTLgenproof.v | 12 ++++++------ src/hls/RTLBlockInstr.v | 6 +++--- 3 files changed, 18 insertions(+), 18 deletions(-) diff --git a/src/Compiler.v b/src/Compiler.v index 7aef71c..8882686 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -68,7 +68,7 @@ Require Import vericert.hls.HTLgenproof. (** ** Declarations We then need to declare the external OCaml functions used to print out intermediate steps in the -compilation, such as ``print_RTL``, ``print_HTL`` and ``print_RTLBlock``. *) +compilation, such as [print_RTL], [print_HTL] and [print_RTLBlock]. *) Parameter print_RTL: Z -> RTL.program -> unit. Parameter print_HTL: Z -> HTL.program -> unit. @@ -158,7 +158,7 @@ Qed. (** *** Top-level Translation -An optimised transformation function from ``RTL`` to ``Verilog`` can then be defined, which applies +An optimised transformation function from [RTL] to [Verilog] can then be defined, which applies the front end compiler optimisations of CompCert to the RTL that is generated and then performs the two Vericert passes from RTL to HTL and then from HTL to Verilog. *) @@ -236,7 +236,7 @@ Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := (** ** Correctness Proof -Finally, the top-level definition for all the passes is defined, which combines the ``match_prog`` +Finally, the top-level definition for all the passes is defined, which combines the [match_prog] predicates of each translation pass from C until Verilog. *) Local Open Scope linking_scope. @@ -260,14 +260,14 @@ Definition CompCert's_passes := ::: mkpass Veriloggenproof.match_prog ::: pass_nil _. -(** These passes are then composed into a larger, top-level ``match_prog`` predicate which matches a +(** These passes are then composed into a larger, top-level [match_prog] predicate which matches a C program directly with a Verilog program. *) Definition match_prog: Csyntax.program -> Verilog.program -> Prop := pass_match (compose_passes CompCert's_passes). (** We then need to prove that this predicate holds, assuming that the translation is performed -using the ``transf_hls`` function declared above. *) +using the [transf_hls] function declared above. *) Theorem transf_hls_match: forall p tp, @@ -373,8 +373,8 @@ Qed. The following theorem is a *backward simulation* between the C and Verilog, which proves the semantics preservation of the translation. We can assume that the C and Verilog programs match, as -we have proven previously in ``transf_hls_match`` that our translation implies that the -``match_prog`` predicate will hold. *) +we have proven previously in [transf_hls_match] that our translation implies that the +[match_prog] predicate will hold. *) Theorem c_semantic_preservation: forall p tp, @@ -391,8 +391,8 @@ Proof. exact (proj2 (cstrategy_semantic_preservation _ _ H)). Qed. -(** We can then use ``transf_hls_match`` to prove the backward simulation where the assumption is -that the translation is performed using the ``transf_hls`` function and that it succeeds. *) +(** We can then use [transf_hls_match] to prove the backward simulation where the assumption is +that the translation is performed using the [transf_hls] function and that it succeeds. *) Theorem transf_c_program_correct: forall p tp, 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 7dabf0b931036148fd7986d7d75624d81c07f146 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 27 Feb 2022 12:03:52 +0000 Subject: Update documentation generation --- docs/Makefile | 31 ++ docs/documentation.org | 13 +- docs/images/toolflow.pdf | Bin 0 -> 20490 bytes docs/images/toolflow.png | Bin 0 -> 26864 bytes docs/images/toolflow.svg | 1250 ++++++++++++++++++++++++++++++++++++++++++++ docs/man.org | 2 + docs/res/install-deps.el | 10 + docs/res/publish-html.el | 12 + docs/res/publish-man.el | 12 + docs/res/publish-manual.el | 12 + 10 files changed, 1336 insertions(+), 6 deletions(-) create mode 100644 docs/Makefile create mode 100644 docs/images/toolflow.pdf create mode 100644 docs/images/toolflow.png create mode 100644 docs/images/toolflow.svg create mode 100644 docs/res/install-deps.el create mode 100644 docs/res/publish-html.el create mode 100644 docs/res/publish-man.el create mode 100644 docs/res/publish-manual.el diff --git a/docs/Makefile b/docs/Makefile new file mode 100644 index 0000000..93083cd --- /dev/null +++ b/docs/Makefile @@ -0,0 +1,31 @@ +all: manual src man-html + +install-deps: + emacs --batch --load ./res/install-deps.el + +manual: + mkdir -p manual + emacs --batch --file documentation.org --load ./res/publish-manual.el + makeinfo --html --number-sections --no-split \ + --css-ref "https://www.gnu.org/software/emacs/manual.css" \ + vericert.texi -o ./manual/index.html + cp -r images ./manual/. + +man-html: + mkdir -p man + emacs --batch --file man.org --load ./res/publish-html.el + cp man.html ./man/vericert.1.html + +src: + $(MAKE) -C .. doc + cp -r ../html src + +upload: + tar caf docs.tar.xz manual man src + rsync docs.tar.xz "zk@leika.ymhg.org:~" + +.PHONY: all upload manual man src install-deps man-html + +clean: + rm -rf manual man src + rm -f docs.tar.xz diff --git a/docs/documentation.org b/docs/documentation.org index 3a0633f..e951f83 100644 --- a/docs/documentation.org +++ b/docs/documentation.org @@ -4,7 +4,7 @@ #+email: git@ymhg.org #+language: en -* Docs +* Introduction :PROPERTIES: :EXPORT_FILE_NAME: _index :EXPORT_HUGO_SECTION: docs @@ -15,10 +15,10 @@ Vericert translates C code into a hardware description language called Verilog, synthesised into hardware, to be placed onto a field-programmable gate array (FPGA) or application-specific integrated circuit (ASIC). -#+attr_html: :width 600 +#+attr_html: :width "600px" #+caption: Current design of Vericert, where HTL is an intermediate language representing a finite state machine with data-path (FSMD) and Verilog is the target language. #+name: fig:design -[[./images/toolflow.svg]] +[[./images/toolflow.png]] The design shown in Figure [[fig:design]] shows how Vericert leverages an existing verified C compiler called [[https://compcert.org/compcert-C.html][CompCert]] to perform this translation. @@ -49,9 +49,9 @@ Free Documentation License''. :END: #+transclude: [[file:~/projects/vericert/README.org::#building][file:~/projects/vericert/README.org::#building]] :only-contents :exclude-elements "headline property-drawer" -#+transclude: [[file:~/projects/vericert/README.org::#downloading-compcert][file:~/projects/vericert/README.org::#downloading-compcert]] -#+transclude: [[file:~/projects/vericert/README.org::#setting-up-nix][file:~/projects/vericert/README.org::#setting-up-nix]] -#+transclude: [[file:~/projects/vericert/README.org::#makefile-build][file:~/projects/vericert/README.org::#makefile-build]] +#+transclude: [[file:~/projects/vericert/README.org::#downloading-compcert][file:~/projects/vericert/README.org::#downloading-compcert]] :level 2 +#+transclude: [[file:~/projects/vericert/README.org::#setting-up-nix][file:~/projects/vericert/README.org::#setting-up-nix]] :level 2 +#+transclude: [[file:~/projects/vericert/README.org::#makefile-build][file:~/projects/vericert/README.org::#makefile-build]] :level 2 ** Testing @@ -155,6 +155,7 @@ to the proper documentation. :PROPERTIES: :CUSTOM_ID: scheduling :END: +#+cindex: scheduling Scheduling is an optimisation which is used to run various instructions in parallel that are independent to each other. diff --git a/docs/images/toolflow.pdf b/docs/images/toolflow.pdf new file mode 100644 index 0000000..5fcee67 Binary files /dev/null and b/docs/images/toolflow.pdf differ diff --git a/docs/images/toolflow.png b/docs/images/toolflow.png new file mode 100644 index 0000000..601d6be Binary files /dev/null and b/docs/images/toolflow.png differ diff --git a/docs/images/toolflow.svg b/docs/images/toolflow.svg new file mode 100644 index 0000000..0d8f39f --- /dev/null +++ b/docs/images/toolflow.svg @@ -0,0 +1,1250 @@ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + diff --git a/docs/man.org b/docs/man.org index 36dbc42..f4d8aef 100644 --- a/docs/man.org +++ b/docs/man.org @@ -1,6 +1,8 @@ #+title: vericert #+man_class_options: :section-id "1" #+export_file_name: vericert.1 +#+options: toc:nil num:nil +#+html_head_extra: * NAME diff --git a/docs/res/install-deps.el b/docs/res/install-deps.el new file mode 100644 index 0000000..9440db5 --- /dev/null +++ b/docs/res/install-deps.el @@ -0,0 +1,10 @@ +(require 'package) +(package-initialize) +(add-to-list 'package-archives '("nongnu" . "https://elpa.nongnu.org/nongnu/") t) +(add-to-list 'package-archives '("melpa" . "https://melpa.org/packages/") t) +(add-to-list 'package-archives '("gnu-devel" . "https://elpa.gnu.org/devel/") t) +(package-refresh-contents) + +(package-install 'org) +(package-install 'org-contrib) +(package-install 'org-transclusion) diff --git a/docs/res/publish-html.el b/docs/res/publish-html.el new file mode 100644 index 0000000..68de36d --- /dev/null +++ b/docs/res/publish-html.el @@ -0,0 +1,12 @@ +(require 'package) +(package-initialize) + +(require 'org) +(require 'org-transclusion) +(require 'ox) +(require 'ox-html) + +(setq org-transclusion-exclude-elements nil) + +(org-transclusion-add-all) +(org-html-export-to-html) diff --git a/docs/res/publish-man.el b/docs/res/publish-man.el new file mode 100644 index 0000000..634b454 --- /dev/null +++ b/docs/res/publish-man.el @@ -0,0 +1,12 @@ +(require 'package) +(package-initialize) + +(require 'org) +(require 'org-transclusion) +(require 'ox) +(require 'ox-man) + +(setq org-transclusion-exclude-elements nil) + +(org-transclusion-add-all) +(org-man-export-to-man) diff --git a/docs/res/publish-manual.el b/docs/res/publish-manual.el new file mode 100644 index 0000000..df5db6b --- /dev/null +++ b/docs/res/publish-manual.el @@ -0,0 +1,12 @@ +(require 'package) +(package-initialize) + +(require 'org) +(require 'org-transclusion) +(require 'ox) +(require 'ox-texinfo) + +(setq org-transclusion-exclude-elements nil) + +(org-transclusion-add-all) +(org-texinfo-export-to-texinfo) -- cgit From 8d8a3a544682a3689e99533ede019cd60a8e9ce6 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 27 Feb 2022 16:35:21 +0000 Subject: Update documentation links --- docs/documentation.org | 12 ++++++++---- docs/man.org | 11 +++++------ docs/res/install-deps.el | 1 + docs/res/publish-html.el | 12 +++++++++++- 4 files changed, 25 insertions(+), 11 deletions(-) diff --git a/docs/documentation.org b/docs/documentation.org index e951f83..85bc8fd 100644 --- a/docs/documentation.org +++ b/docs/documentation.org @@ -48,10 +48,10 @@ Free Documentation License''. :CUSTOM_ID: building :END: -#+transclude: [[file:~/projects/vericert/README.org::#building][file:~/projects/vericert/README.org::#building]] :only-contents :exclude-elements "headline property-drawer" -#+transclude: [[file:~/projects/vericert/README.org::#downloading-compcert][file:~/projects/vericert/README.org::#downloading-compcert]] :level 2 -#+transclude: [[file:~/projects/vericert/README.org::#setting-up-nix][file:~/projects/vericert/README.org::#setting-up-nix]] :level 2 -#+transclude: [[file:~/projects/vericert/README.org::#makefile-build][file:~/projects/vericert/README.org::#makefile-build]] :level 2 +#+transclude: [[file:~/projects/vericert/README.org::#building][file:../README.org::#building]] :only-contents :exclude-elements "headline property-drawer" +#+transclude: [[file:~/projects/vericert/README.org::#downloading-compcert][file:../README.org::#downloading-compcert]] :level 2 +#+transclude: [[file:~/projects/vericert/README.org::#setting-up-nix][file:../README.org::#setting-up-nix]] :level 2 +#+transclude: [[file:~/projects/vericert/README.org::#makefile-build][file:../README.org::#makefile-build]] :level 2 ** Testing @@ -132,6 +132,10 @@ $ echo $? 50 #+end_src +** Man pages + +#+transclude: [[file:man.org][file:man.org]] :exclude-elements "keyword" :level 3 + * Unreleased Features :PROPERTIES: :EXPORT_FILE_NAME: unreleased diff --git a/docs/man.org b/docs/man.org index f4d8aef..cb6143f 100644 --- a/docs/man.org +++ b/docs/man.org @@ -1,6 +1,5 @@ #+title: vericert #+man_class_options: :section-id "1" -#+export_file_name: vericert.1 #+options: toc:nil num:nil #+html_head_extra: @@ -14,19 +13,19 @@ vericert - A formally verified high-level synthesis tool. * DESCRIPTION -** HLS Options +** HLS Options: - --no-hls :: Do not use HLS and generate standard flow - --simulate :: Simulate the result with the Verilog semantics - --debug-hls :: Add debug logic to the Verilog - --initialise-stack :: initialise the stack to all 0s -** HLS Optimisations +** HLS Optimisations: - -fschedule :: Schedule the resulting hardware [off] - -fif-conversion :: If-conversion optimisation [off] -** General options +** General options: - -stdlib :: Set the path of the Compcert run-time library - -v :: Print external commands before invoking them @@ -36,7 +35,7 @@ vericert - A formally verified high-level synthesis tool. - -conf :: Read configuration from file - @ :: Read command line options from -** Tracing Options +** Tracing Options: - -dprepro :: Save C file after preprocessing in .i - -dparse :: Save C file after parsing and elaboration in .parsed.c @@ -53,7 +52,7 @@ vericert - A formally verified high-level synthesis tool. - -sdump :: Save info for post-linking validation in .json - -o :: Generate output in -** Diagnostic options +** Diagnostic options: - -Wall :: Enable all warnings - -W :: Enable the specific diff --git a/docs/res/install-deps.el b/docs/res/install-deps.el index 9440db5..09cf4ae 100644 --- a/docs/res/install-deps.el +++ b/docs/res/install-deps.el @@ -8,3 +8,4 @@ (package-install 'org) (package-install 'org-contrib) (package-install 'org-transclusion) +(package-install 'htmlize) diff --git a/docs/res/publish-html.el b/docs/res/publish-html.el index 68de36d..a9c7a9d 100644 --- a/docs/res/publish-html.el +++ b/docs/res/publish-html.el @@ -5,8 +5,18 @@ (require 'org-transclusion) (require 'ox) (require 'ox-html) +(require 'htmlize) -(setq org-transclusion-exclude-elements nil) +(setq org-transclusion-exclude-elements nil + org-html-head-include-default-style nil + org-html-head-include-scripts nil + org-html-postamble-format '(("en" "")) + org-html-postamble t + org-html-divs '((preamble "header" "header") + (content "article" "content") + (postamble "footer" "postamble")) + org-html-doctype "html5" + org-html-htmlize-output-type 'css) (org-transclusion-add-all) (org-html-export-to-html) -- cgit From 71669aaba76f65ff9ac99a99f4559fdfcbf3446d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 27 Feb 2022 18:13:15 +0000 Subject: Remove .envrc as it should be local --- .envrc | 1 - 1 file changed, 1 deletion(-) delete mode 100644 .envrc diff --git a/.envrc b/.envrc deleted file mode 100644 index 1d953f4..0000000 --- a/.envrc +++ /dev/null @@ -1 +0,0 @@ -use nix -- cgit From 7a09bbe8d18abb376104e1399fc84ce911d8897d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 28 Feb 2022 09:41:13 +0000 Subject: Add ptsets.ml library from CompCert for Bourdoncle --- src/bourdoncle/ptset.ml | 666 +++++++++++++++++++++++++++++++++++++++++++++++ src/bourdoncle/ptset.mli | 125 +++++++++ 2 files changed, 791 insertions(+) create mode 100644 src/bourdoncle/ptset.ml create mode 100644 src/bourdoncle/ptset.mli diff --git a/src/bourdoncle/ptset.ml b/src/bourdoncle/ptset.ml new file mode 100644 index 0000000..c94691b --- /dev/null +++ b/src/bourdoncle/ptset.ml @@ -0,0 +1,666 @@ +(* + * Copyright (C) 2008 Jean-Christophe Filliatre + * Copyright (C) 2008, 2009 Laurent Hubert (CNRS) + * + * This software is free software; you can redistribute it and/or + * modify it under the terms of the GNU Lesser General Public License + * version 2.1, with the special exception on linking described in file + * LICENSE. + + * This software 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. + *) + + +(*s Sets of integers implemented as Patricia trees, following Chris + Okasaki and Andrew Gill's paper {\em Fast Mergeable Integer Maps} + ({\tt\small http://www.cs.columbia.edu/\~{}cdo/papers.html\#ml98maps}). + Patricia trees provide faster operations than standard library's + module [Set], and especially very fast [union], [subset], [inter] + and [diff] operations. *) + +(*s The idea behind Patricia trees is to build a {\em trie} on the + binary digits of the elements, and to compact the representation + by branching only one the relevant bits (i.e. the ones for which + there is at least on element in each subtree). We implement here + {\em little-endian} Patricia trees: bits are processed from + least-significant to most-significant. The trie is implemented by + the following type [t]. [Empty] stands for the empty trie, and + [Leaf k] for the singleton [k]. (Note that [k] is the actual + element.) [Branch (m,p,l,r)] represents a branching, where [p] is + the prefix (from the root of the trie) and [m] is the branching + bit (a power of 2). [l] and [r] contain the subsets for which the + branching bit is respectively 0 and 1. Invariant: the trees [l] + and [r] are not empty. *) + + +module type S = sig + type t + type elt = int + val empty : t + val is_empty : t -> bool + val mem : int -> t -> bool + val add : int -> t -> t + val singleton : int -> t + val remove : int -> t -> t + val union : t -> t -> t + val subset : t -> t -> bool + val inter : t -> t -> t + val diff : t -> t -> t + val equal : t -> t -> bool + val compare : t -> t -> int + val elements : t -> int list + val choose : t -> int + val cardinal : t -> int + val iter : (int -> unit) -> t -> unit + val fold : (int -> 'a -> 'a) -> t -> 'a -> 'a + val for_all : (int -> bool) -> t -> bool + val exists : (int -> bool) -> t -> bool + val filter : (int -> bool) -> t -> t + val partition : (int -> bool) -> t -> t * t + val split : int -> t -> t * bool * t + val min_elt : t -> int + val max_elt : t -> int + val intersect : t -> t -> bool + val choose_and_remove : t -> int*t +end + + +(*i*) +type elt = int +(*i*) + +type t = + | Empty + | Leaf of int + | Branch of int * int * t * t + +(*s Example: the representation of the set $\{1,4,5\}$ is + $$\mathtt{Branch~(0,~1,~Leaf~4,~Branch~(1,~4,~Leaf~1,~Leaf~5))}$$ + The first branching bit is the bit 0 (and the corresponding prefix + is [0b0], not of use here), with $\{4\}$ on the left and $\{1,5\}$ on the + right. Then the right subtree branches on bit 2 (and so has a branching + value of $2^2 = 4$), with prefix [0b01 = 1]. *) + +(*s Empty set and singletons. *) + +let empty = Empty + +let is_empty = function Empty -> true | _ -> false + +let singleton k = Leaf k + +(*s Testing the occurrence of a value is similar to the search in a + binary search tree, where the branching bit is used to select the + appropriate subtree. *) + +let zero_bit k m = (k land m) == 0 + +let rec mem k = function + | Empty -> false + | Leaf j -> k == j + | Branch (_, m, l, r) -> mem k (if zero_bit k m then l else r) + +(*s The following operation [join] will be used in both insertion and + union. Given two non-empty trees [t0] and [t1] with longest common + prefixes [p0] and [p1] respectively, which are supposed to + disagree, it creates the union of [t0] and [t1]. For this, it + computes the first bit [m] where [p0] and [p1] disagree and create + a branching node on that bit. Depending on the value of that bit + in [p0], [t0] will be the left subtree and [t1] the right one, or + the converse. Computing the first branching bit of [p0] and [p1] + uses a nice property of twos-complement representation of integers. *) + +let lowest_bit x = x land (-x) + +let branching_bit p0 p1 = lowest_bit (p0 lxor p1) + +let mask p m = p land (m-1) + +let join (p0,t0,p1,t1) = + let m = branching_bit p0 p1 in + if zero_bit p0 m then + Branch (mask p0 m, m, t0, t1) + else + Branch (mask p0 m, m, t1, t0) + +(*s Then the insertion of value [k] in set [t] is easily implemented + using [join]. Insertion in a singleton is just the identity or a + call to [join], depending on the value of [k]. When inserting in + a branching tree, we first check if the value to insert [k] + matches the prefix [p]: if not, [join] will take care of creating + the above branching; if so, we just insert [k] in the appropriate + subtree, depending of the branching bit. *) + +let match_prefix k p m = (mask k m) == p + +let add k t = + let rec ins = function + | Empty -> Leaf k + | Leaf j as t -> + if j == k then t else join (k, Leaf k, j, t) + | Branch (p,m,t0,t1) as t -> + if match_prefix k p m then + if zero_bit k m then + Branch (p, m, ins t0, t1) + else + Branch (p, m, t0, ins t1) + else + join (k, Leaf k, p, t) + in + ins t + +(*s The code to remove an element is basically similar to the code of + insertion. But since we have to maintain the invariant that both + subtrees of a [Branch] node are non-empty, we use here the + ``smart constructor'' [branch] instead of [Branch]. *) + +let branch = function + | (_,_,Empty,t) -> t + | (_,_,t,Empty) -> t + | (p,m,t0,t1) -> Branch (p,m,t0,t1) + +let remove k t = + let rec rmv = function + | Empty -> Empty + | Leaf j as t -> if k == j then Empty else t + | Branch (p,m,t0,t1) as t -> + if match_prefix k p m then + if zero_bit k m then + branch (p, m, rmv t0, t1) + else + branch (p, m, t0, rmv t1) + else + t + in + rmv t + +let rec choose_and_remove = function + | Empty -> raise Not_found + | Leaf j -> (j,Empty) + | Branch (p,m,t0,t1) -> + let (j,t0') = choose_and_remove t0 + in (j, branch (p,m,t0',t1)) + +(*s One nice property of Patricia trees is to support a fast union + operation (and also fast subset, difference and intersection + operations). When merging two branching trees we examine the + following four cases: (1) the trees have exactly the same + prefix; (2/3) one prefix contains the other one; and (4) the + prefixes disagree. In cases (1), (2) and (3) the recursion is + immediate; in case (4) the function [join] creates the appropriate + branching. *) + +let rec merge = function + | t1,t2 when t1==t2 -> t1 + | Empty, t -> t + | t, Empty -> t + | Leaf k, t -> add k t + | t, Leaf k -> add k t + | (Branch (p,m,s0,s1) as s), (Branch (q,n,t0,t1) as t) -> + if m == n && match_prefix q p m then + (* The trees have the same prefix. Merge the subtrees. *) + Branch (p, m, merge (s0,t0), merge (s1,t1)) + else if m < n && match_prefix q p m then + (* [q] contains [p]. Merge [t] with a subtree of [s]. *) + if zero_bit q m then + Branch (p, m, merge (s0,t), s1) + else + Branch (p, m, s0, merge (s1,t)) + else if m > n && match_prefix p q n then + (* [p] contains [q]. Merge [s] with a subtree of [t]. *) + if zero_bit p n then + Branch (q, n, merge (s,t0), t1) + else + Branch (q, n, t0, merge (s,t1)) + else + (* The prefixes disagree. *) + join (p, s, q, t) + +let union s t = merge (s,t) + +(*s When checking if [s1] is a subset of [s2] only two of the above + four cases are relevant: when the prefixes are the same and when the + prefix of [s1] contains the one of [s2], and then the recursion is + obvious. In the other two cases, the result is [false]. *) + +let rec subset s1 s2 = match (s1,s2) with + | Empty, _ -> true + | _, Empty -> false + | Leaf k1, _ -> mem k1 s2 + | Branch _, Leaf _ -> false + | Branch (p1,m1,l1,r1), Branch (p2,m2,l2,r2) -> + if m1 == m2 && p1 == p2 then + subset l1 l2 && subset r1 r2 + else if m1 > m2 && match_prefix p1 p2 m2 then + if zero_bit p1 m2 then + subset l1 l2 && subset r1 l2 + else + subset l1 r2 && subset r1 r2 + else + false + +(*s To compute the intersection and the difference of two sets, we + still examine the same four cases as in [merge]. The recursion is + then obvious. *) + +let rec inter s1 s2 = match (s1,s2) with + | t1, t2 when t1==t2 -> t1 + | Empty, _ -> Empty + | _, Empty -> Empty + | Leaf k1, _ -> if mem k1 s2 then s1 else Empty + | _, Leaf k2 -> if mem k2 s1 then s2 else Empty + | Branch (p1,m1,l1,r1), Branch (p2,m2,l2,r2) -> + if m1 == m2 && p1 == p2 then + merge (inter l1 l2, inter r1 r2) + else if m1 < m2 && match_prefix p2 p1 m1 then + inter (if zero_bit p2 m1 then l1 else r1) s2 + else if m1 > m2 && match_prefix p1 p2 m2 then + inter s1 (if zero_bit p1 m2 then l2 else r2) + else + Empty + +let rec diff s1 s2 = match (s1,s2) with + | Empty, _ -> Empty + | _, Empty -> s1 + | Leaf k1, _ -> if mem k1 s2 then Empty else s1 + | _, Leaf k2 -> remove k2 s1 + | Branch (p1,m1,l1,r1), Branch (p2,m2,l2,r2) -> + if m1 == m2 && p1 == p2 then + merge (diff l1 l2, diff r1 r2) + else if m1 < m2 && match_prefix p2 p1 m1 then + if zero_bit p2 m1 then + merge (diff l1 s2, r1) + else + merge (l1, diff r1 s2) + else if m1 > m2 && match_prefix p1 p2 m2 then + if zero_bit p1 m2 then diff s1 l2 else diff s1 r2 + else + s1 + +(*s All the following operations ([cardinal], [iter], [fold], [for_all], + [exists], [filter], [partition], [choose], [elements]) are + implemented as for any other kind of binary trees. *) + +let rec cardinal = function + | Empty -> 0 + | Leaf _ -> 1 + | Branch (_,_,t0,t1) -> cardinal t0 + cardinal t1 + +let rec iter f = function + | Empty -> () + | Leaf k -> f k + | Branch (_,_,t0,t1) -> iter f t0; iter f t1 + +let rec fold f s accu = match s with + | Empty -> accu + | Leaf k -> f k accu + | Branch (_,_,t0,t1) -> fold f t0 (fold f t1 accu) + +let rec for_all p = function + | Empty -> true + | Leaf k -> p k + | Branch (_,_,t0,t1) -> for_all p t0 && for_all p t1 + +let rec exists p = function + | Empty -> false + | Leaf k -> p k + | Branch (_,_,t0,t1) -> exists p t0 || exists p t1 + +let rec filter pr = function + | Empty -> Empty + | Leaf k as t -> if pr k then t else Empty + | Branch (p,m,t0,t1) -> branch (p, m, filter pr t0, filter pr t1) + +let partition p s = + let rec part (t,f as acc) = function + | Empty -> acc + | Leaf k -> if p k then (add k t, f) else (t, add k f) + | Branch (_,_,t0,t1) -> part (part acc t0) t1 + in + part (Empty, Empty) s + +let rec choose = function + | Empty -> raise Not_found + | Leaf k -> k + | Branch (_, _,t0,_) -> choose t0 (* we know that [t0] is non-empty *) + +let elements s = + let rec elements_aux acc = function + | Empty -> acc + | Leaf k -> k :: acc + | Branch (_,_,l,r) -> elements_aux (elements_aux acc l) r + in + (* unfortunately there is no easy way to get the elements in ascending + order with little-endian Patricia trees *) + List.sort Stdlib.compare (elements_aux [] s) + +let split x s = + let coll k (l, b, r) = + if k < x then add k l, b, r + else if k > x then l, b, add k r + else l, true, r + in + fold coll s (Empty, false, Empty) + +(*s There is no way to give an efficient implementation of [min_elt] + and [max_elt], as with binary search trees. The following + implementation is a traversal of all elements, barely more + efficient than [fold min t (choose t)] (resp. [fold max t (choose + t)]). Note that we use the fact that there is no constructor + [Empty] under [Branch] and therefore always a minimal + (resp. maximal) element there. *) + +let rec min_elt = function + | Empty -> raise Not_found + | Leaf k -> k + | Branch (_,_,s,t) -> min (min_elt s) (min_elt t) + +let rec max_elt = function + | Empty -> raise Not_found + | Leaf k -> k + | Branch (_,_,s,t) -> max (max_elt s) (max_elt t) + +(*s Another nice property of Patricia trees is to be independent of the + order of insertion. As a consequence, two Patricia trees have the + same elements if and only if they are structurally equal. *) + +let equal v1 v2 = 0 == (compare v1 v2) + +let compare = compare + +(*i*) +(* let make l = List.fold_right add l empty *) +(*i*) + +(*s Additional functions w.r.t to [Set.S]. *) + +let rec intersect s1 s2 = match (s1,s2) with + | Empty, _ -> false + | _, Empty -> false + | Leaf k1, _ -> mem k1 s2 + | _, Leaf k2 -> mem k2 s1 + | Branch (p1,m1,l1,r1), Branch (p2,m2,l2,r2) -> + if m1 == m2 && p1 == p2 then + intersect l1 l2 || intersect r1 r2 + else if m1 < m2 && match_prefix p2 p1 m1 then + intersect (if zero_bit p2 m1 then l1 else r1) s2 + else if m1 > m2 && match_prefix p1 p2 m2 then + intersect s1 (if zero_bit p1 m2 then l2 else r2) + else + false + + +(*s Big-endian Patricia trees *) + +module Big = struct + + type elt = int + + type t_ = t + type t = t_ + + let empty = Empty + + let is_empty = function Empty -> true | _ -> false + + let singleton k = Leaf k + + let zero_bit k m = (k land m) == 0 + + let rec mem k = function + | Empty -> false + | Leaf j -> k == j + | Branch (_, m, l, r) -> mem k (if zero_bit k m then l else r) + + let mask k m = (k lor (m-1)) land (lnot m) + + (* we first write a naive implementation of [highest_bit] + only has to work for bytes *) + let naive_highest_bit x = + assert (x < 256); + let rec loop i = + if i = 0 then 1 else if x lsr i = 1 then 1 lsl i else loop (i-1) + in + loop 7 + + (* then we build a table giving the highest bit for bytes *) + let hbit = Array.init 256 naive_highest_bit + + (* to determine the highest bit of [x] we split it into bytes *) + let highest_bit_32 x = + let n = x lsr 24 in if n != 0 then hbit.(n) lsl 24 + else let n = x lsr 16 in if n != 0 then hbit.(n) lsl 16 + else let n = x lsr 8 in if n != 0 then hbit.(n) lsl 8 + else hbit.(x) + + let highest_bit_64 x = + let n = x lsr 32 in if n != 0 then (highest_bit_32 n) lsl 32 + else highest_bit_32 x + + let highest_bit = match Sys.word_size with + | 32 -> highest_bit_32 + | 64 -> highest_bit_64 + | _ -> assert false + + let branching_bit p0 p1 = highest_bit (p0 lxor p1) + + let join (p0,t0,p1,t1) = + (*i let m = function Branch (_,m,_,_) -> m | _ -> 0 in i*) + let m = branching_bit p0 p1 (*EXP (m t0) (m t1) *) in + if zero_bit p0 m then + Branch (mask p0 m, m, t0, t1) + else + Branch (mask p0 m, m, t1, t0) + + let match_prefix k p m = (mask k m) == p + + let add k t = + let rec ins = function + | Empty -> Leaf k + | Leaf j as t -> + if j == k then t else join (k, Leaf k, j, t) + | Branch (p,m,t0,t1) as t -> + if match_prefix k p m then + if zero_bit k m then + Branch (p, m, ins t0, t1) + else + Branch (p, m, t0, ins t1) + else + join (k, Leaf k, p, t) + in + ins t + + let remove k t = + let rec rmv = function + | Empty -> Empty + | Leaf j as t -> if k == j then Empty else t + | Branch (p,m,t0,t1) as t -> + if match_prefix k p m then + if zero_bit k m then + branch (p, m, rmv t0, t1) + else + branch (p, m, t0, rmv t1) + else + t + in + rmv t + + let rec merge = function + | Empty, t -> t + | t, Empty -> t + | Leaf k, t -> add k t + | t, Leaf k -> add k t + | (Branch (p,m,s0,s1) as s), (Branch (q,n,t0,t1) as t) -> + if m == n && match_prefix q p m then + (* The trees have the same prefix. Merge the subtrees. *) + Branch (p, m, merge (s0,t0), merge (s1,t1)) + else if m > n && match_prefix q p m then + (* [q] contains [p]. Merge [t] with a subtree of [s]. *) + if zero_bit q m then + Branch (p, m, merge (s0,t), s1) + else + Branch (p, m, s0, merge (s1,t)) + else if m < n && match_prefix p q n then + (* [p] contains [q]. Merge [s] with a subtree of [t]. *) + if zero_bit p n then + Branch (q, n, merge (s,t0), t1) + else + Branch (q, n, t0, merge (s,t1)) + else + (* The prefixes disagree. *) + join (p, s, q, t) + + let union s t = merge (s,t) + + let rec subset s1 s2 = match (s1,s2) with + | Empty, _ -> true + | _, Empty -> false + | Leaf k1, _ -> mem k1 s2 + | Branch _, Leaf _ -> false + | Branch (p1,m1,l1,r1), Branch (p2,m2,l2,r2) -> + if m1 == m2 && p1 == p2 then + subset l1 l2 && subset r1 r2 + else if m1 < m2 && match_prefix p1 p2 m2 then + if zero_bit p1 m2 then + subset l1 l2 && subset r1 l2 + else + subset l1 r2 && subset r1 r2 + else + false + + let rec inter s1 s2 = match (s1,s2) with + | Empty, _ -> Empty + | _, Empty -> Empty + | Leaf k1, _ -> if mem k1 s2 then s1 else Empty + | _, Leaf k2 -> if mem k2 s1 then s2 else Empty + | Branch (p1,m1,l1,r1), Branch (p2,m2,l2,r2) -> + if m1 == m2 && p1 == p2 then + merge (inter l1 l2, inter r1 r2) + else if m1 > m2 && match_prefix p2 p1 m1 then + inter (if zero_bit p2 m1 then l1 else r1) s2 + else if m1 < m2 && match_prefix p1 p2 m2 then + inter s1 (if zero_bit p1 m2 then l2 else r2) + else + Empty + + let rec diff s1 s2 = match (s1,s2) with + | Empty, _ -> Empty + | _, Empty -> s1 + | Leaf k1, _ -> if mem k1 s2 then Empty else s1 + | _, Leaf k2 -> remove k2 s1 + | Branch (p1,m1,l1,r1), Branch (p2,m2,l2,r2) -> + if m1 == m2 && p1 == p2 then + merge (diff l1 l2, diff r1 r2) + else if m1 > m2 && match_prefix p2 p1 m1 then + if zero_bit p2 m1 then + merge (diff l1 s2, r1) + else + merge (l1, diff r1 s2) + else if m1 < m2 && match_prefix p1 p2 m2 then + if zero_bit p1 m2 then diff s1 l2 else diff s1 r2 + else + s1 + + (* same implementation as for little-endian Patricia trees *) + let cardinal = cardinal + let iter = iter + let fold = fold + let for_all = for_all + let exists = exists + let filter = filter + + let partition p s = + let rec part (t,f as acc) = function + | Empty -> acc + | Leaf k -> if p k then (add k t, f) else (t, add k f) + | Branch (_,_,t0,t1) -> part (part acc t0) t1 + in + part (Empty, Empty) s + + let choose = choose + let choose_and_remove = choose_and_remove + + let elements s = + let rec elements_aux acc = function + | Empty -> acc + | Leaf k -> k :: acc + | Branch (_,_,l,r) -> elements_aux (elements_aux acc r) l + in + (* we still have to sort because of possible negative elements *) + List.sort Stdlib.compare (elements_aux [] s) + + let split x s = + let coll k (l, b, r) = + if k < x then add k l, b, r + else if k > x then l, b, add k r + else l, true, r + in + fold coll s (Empty, false, Empty) + + (* could be slightly improved (when we now that a branch contains only + positive or only negative integers) *) + let min_elt = min_elt + let max_elt = max_elt + + let equal = (=) + + let compare = compare + + (* let make l = List.fold_right add l empty *) + + let rec intersect s1 s2 = match (s1,s2) with + | Empty, _ -> false + | _, Empty -> false + | Leaf k1, _ -> mem k1 s2 + | _, Leaf k2 -> mem k2 s1 + | Branch (p1,m1,l1,r1), Branch (p2,m2,l2,r2) -> + if m1 == m2 && p1 == p2 then + intersect l1 l2 || intersect r1 r2 + else if m1 > m2 && match_prefix p2 p1 m1 then + intersect (if zero_bit p2 m1 then l1 else r1) s2 + else if m1 < m2 && match_prefix p1 p2 m2 then + intersect s1 (if zero_bit p1 m2 then l2 else r2) + else + false + +end + +(*s Big-endian Patricia trees with non-negative elements only *) + +module BigPos = struct + + include Big + + let singleton x = if x < 0 then invalid_arg "BigPos.singleton"; singleton x + + let add x s = if x < 0 then invalid_arg "BigPos.add"; add x s + + (* Patricia trees are now binary search trees! *) + + let rec mem k = function + | Empty -> false + | Leaf j -> k == j + | Branch (p, _, l, r) -> if k <= p then mem k l else mem k r + + let rec min_elt = function + | Empty -> raise Not_found + | Leaf k -> k + | Branch (_,_,s,_) -> min_elt s + + let rec max_elt = function + | Empty -> raise Not_found + | Leaf k -> k + | Branch (_,_,_,t) -> max_elt t + + (* we do not have to sort anymore *) + let elements s = + let rec elements_aux acc = function + | Empty -> acc + | Leaf k -> k :: acc + | Branch (_,_,l,r) -> elements_aux (elements_aux acc r) l + in + elements_aux [] s + +end diff --git a/src/bourdoncle/ptset.mli b/src/bourdoncle/ptset.mli new file mode 100644 index 0000000..88319ea --- /dev/null +++ b/src/bourdoncle/ptset.mli @@ -0,0 +1,125 @@ +(* + * Copyright (C) 2008 Jean-Christophe Filliatre + * Copyright (C) 2008, 2009 Laurent Hubert (CNRS) + * + * This software is free software; you can redistribute it and/or + * modify it under the terms of the GNU Lesser General Public License + * version 2.1, with the special exception on linking described in file + * LICENSE. + + * This software 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. + *) + +(** Sets of integers implemented as Patricia trees. + + This implementation follows Chris Okasaki and Andrew Gill's paper + {e Fast Mergeable Integer Maps}. + + Patricia trees provide faster operations than standard library's + module [Set], and especially very fast [union], [subset], [inter] + and [diff] operations. *) + +(** The idea behind Patricia trees is to build a {e trie} on the + binary digits of the elements, and to compact the representation + by branching only one the relevant bits (i.e. the ones for which + there is at least on element in each subtree). We implement here + {e little-endian} Patricia trees: bits are processed from + least-significant to most-significant. The trie is implemented by + the following type [t]. [Empty] stands for the empty trie, and + [Leaf k] for the singleton [k]. (Note that [k] is the actual + element.) [Branch (m,p,l,r)] represents a branching, where [p] is + the prefix (from the root of the trie) and [m] is the branching + bit (a power of 2). [l] and [r] contain the subsets for which the + branching bit is respectively 0 and 1. Invariant: the trees [l] + and [r] are not empty. *) + +(** The docuemntation is given for function that differs from [Set.S + with type elt = int]. *) + +module type S = sig + type t + + type elt = int + + val empty : t + + val is_empty : t -> bool + + val mem : int -> t -> bool + + val add : int -> t -> t + + val singleton : int -> t + + val remove : int -> t -> t + + val union : t -> t -> t + + val subset : t -> t -> bool + + val inter : t -> t -> t + + val diff : t -> t -> t + + val equal : t -> t -> bool + + val compare : t -> t -> int + + val elements : t -> int list + + val choose : t -> int + + val cardinal : t -> int + + val iter : (int -> unit) -> t -> unit + + val fold : (int -> 'a -> 'a) -> t -> 'a -> 'a + + val for_all : (int -> bool) -> t -> bool + + val exists : (int -> bool) -> t -> bool + + val filter : (int -> bool) -> t -> t + + val partition : (int -> bool) -> t -> t * t + + val split : int -> t -> t * bool * t + + (** Warning: [min_elt] and [max_elt] are linear w.r.t. the size of the + set. In other words, [min_elt t] is barely more efficient than [fold + min t (choose t)]. *) + + val min_elt : t -> int + val max_elt : t -> int + + (** Additional functions not appearing in the signature [Set.S] from + ocaml standard library. *) + + (** [intersect u v] determines if sets [u] and [v] have a non-empty + intersection. *) + + val intersect : t -> t -> bool + + (** [choose_and_remove t] is equivalent (but more efficient) to + [(fun t -> let i = choose t in (i,remove i t)) t] + @author Laurent Hubert*) + val choose_and_remove : t -> int*t + +end + +include S + +(** Big-endian Patricia trees *) + +module Big : S + +(** Big-endian Patricia trees with non-negative elements. Changes: + - [add] and [singleton] raise [Invalid_arg] if a negative element is given + - [mem] is slightly faster (the Patricia tree is now a search tree) + - [min_elt] and [max_elt] are now O(log(N)) + - [elements] returns a list with elements in ascending order + *) + +module BigPos : S -- cgit From 1252ed9252abc91dec15a240e2866d5344d3d81f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 28 Feb 2022 09:46:07 +0000 Subject: Update .gitignore --- .gitignore | 74 +++++++------------------------------------------------------- 1 file changed, 8 insertions(+), 66 deletions(-) diff --git a/.gitignore b/.gitignore index aaabd2a..43c9acc 100644 --- a/.gitignore +++ b/.gitignore @@ -19,17 +19,8 @@ *.vo *.vok *.vos +*.cache .coq-native/ -.csdp.cache -.lia.cache -.nia.cache -.nlia.cache -.nra.cache -csdp.cache -lia.cache -nia.cache -nlia.cache -nra.cache Makefile.coq Makefile.coq.conf @@ -71,62 +62,6 @@ benchmarks/**/*.v *.iver *.dot -/benchmarks/polybench-syn/stencils/seidel-2d -/benchmarks/polybench-syn/stencils/jacobi-2d -/benchmarks/polybench-syn/data-mining/covariance -/benchmarks/polybench-syn/linear-algebra/blas/gemm -/benchmarks/polybench-syn/linear-algebra/blas/gemver -/benchmarks/polybench-syn/linear-algebra/blas/gesummv -/benchmarks/polybench-syn/linear-algebra/blas/symm -/benchmarks/polybench-syn/linear-algebra/blas/syr2k -/benchmarks/polybench-syn/linear-algebra/blas/syrk -/benchmarks/polybench-syn/linear-algebra/blas/trmm -/benchmarks/polybench-syn/linear-algebra/kernels/2mm -/benchmarks/polybench-syn/linear-algebra/kernels/3mm -/benchmarks/polybench-syn/linear-algebra/kernels/atas -/benchmarks/polybench-syn/linear-algebra/kernels/bicg -/benchmarks/polybench-syn/linear-algebra/kernels/doitgen -/benchmarks/polybench-syn/linear-algebra/kernels/mvt -/benchmarks/polybench-syn/linear-algebra/solvers/cholesky -/benchmarks/polybench-syn/linear-algebra/solvers/durbin -/benchmarks/polybench-syn/linear-algebra/solvers/lu -/benchmarks/polybench-syn/linear-algebra/solvers/ludcmp -/benchmarks/polybench-syn/linear-algebra/solvers/trisolv -/benchmarks/polybench-syn/medley/floyd-warshall -/benchmarks/polybench-syn/medley/nussinov -/benchmarks/polybench-syn/stencils/adi -/benchmarks/polybench-syn/stencils/fdtd-2d -/benchmarks/polybench-syn/stencils/heat-3d -/benchmarks/polybench-syn/stencils/jacobi-1d - -/benchmarks/polybench-syn-div/stencils/seidel-2d -/benchmarks/polybench-syn-div/stencils/jacobi-2d -/benchmarks/polybench-syn-div/data-mining/covariance -/benchmarks/polybench-syn-div/linear-algebra/blas/gemm -/benchmarks/polybench-syn-div/linear-algebra/blas/gemver -/benchmarks/polybench-syn-div/linear-algebra/blas/gesummv -/benchmarks/polybench-syn-div/linear-algebra/blas/symm -/benchmarks/polybench-syn-div/linear-algebra/blas/syr2k -/benchmarks/polybench-syn-div/linear-algebra/blas/syrk -/benchmarks/polybench-syn-div/linear-algebra/blas/trmm -/benchmarks/polybench-syn-div/linear-algebra/kernels/2mm -/benchmarks/polybench-syn-div/linear-algebra/kernels/3mm -/benchmarks/polybench-syn-div/linear-algebra/kernels/atas -/benchmarks/polybench-syn-div/linear-algebra/kernels/bicg -/benchmarks/polybench-syn-div/linear-algebra/kernels/doitgen -/benchmarks/polybench-syn-div/linear-algebra/kernels/mvt -/benchmarks/polybench-syn-div/linear-algebra/solvers/cholesky -/benchmarks/polybench-syn-div/linear-algebra/solvers/durbin -/benchmarks/polybench-syn-div/linear-algebra/solvers/lu -/benchmarks/polybench-syn-div/linear-algebra/solvers/ludcmp -/benchmarks/polybench-syn-div/linear-algebra/solvers/trisolv -/benchmarks/polybench-syn-div/medley/floyd-warshall -/benchmarks/polybench-syn-div/medley/nussinov -/benchmarks/polybench-syn-div/stencils/adi -/benchmarks/polybench-syn-div/stencils/fdtd-2d -/benchmarks/polybench-syn-div/stencils/heat-3d -/benchmarks/polybench-syn-div/stencils/jacobi-1d - # Test *.check *.txt @@ -140,3 +75,10 @@ obj_dir/ /*.v .direnv/ +.envrc + +creduce_bug_*/ + +/docs/man/ +/docs/manual/ +/docs/src/ -- cgit From 88d015ba178665ee21b282f364ccf047522e3b1c Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 28 Feb 2022 12:49:35 +0000 Subject: Rename IP using systemverilog extension --- ip/altera.sv | 2 + ip/altera.v | 2 - ip/div_signed.sv | 193 +++++++++++++++++++++++++++++++++++++++++++++++++++++ ip/div_signed.v | 193 ----------------------------------------------------- ip/div_unsigned.sv | 146 ++++++++++++++++++++++++++++++++++++++++ ip/div_unsigned.v | 146 ---------------------------------------- 6 files changed, 341 insertions(+), 341 deletions(-) create mode 100644 ip/altera.sv delete mode 100644 ip/altera.v create mode 100644 ip/div_signed.sv delete mode 100644 ip/div_signed.v create mode 100644 ip/div_unsigned.sv delete mode 100644 ip/div_unsigned.v diff --git a/ip/altera.sv b/ip/altera.sv new file mode 100644 index 0000000..3839cbe --- /dev/null +++ b/ip/altera.sv @@ -0,0 +1,2 @@ +module ALTERA_MF_MEMORY_INITIALIZATION; +endmodule diff --git a/ip/altera.v b/ip/altera.v deleted file mode 100644 index 3839cbe..0000000 --- a/ip/altera.v +++ /dev/null @@ -1,2 +0,0 @@ -module ALTERA_MF_MEMORY_INITIALIZATION; -endmodule diff --git a/ip/div_signed.sv b/ip/div_signed.sv new file mode 100644 index 0000000..7aab57c --- /dev/null +++ b/ip/div_signed.sv @@ -0,0 +1,193 @@ +///////////////////////////////////////////////////////////////////// +//// //// +//// Non-restoring unsigned divider //// +//// //// +//// Author: Richard Herveille //// +//// richard@asics.ws //// +//// www.asics.ws //// +//// //// +///////////////////////////////////////////////////////////////////// +//// //// +//// Copyright (C) 2002 Richard Herveille //// +//// richard@asics.ws //// +//// //// +//// This source file may be used and distributed without //// +//// restriction provided that this copyright statement is not //// +//// removed from the file and that any derivative work contains //// +//// the original copyright notice and the associated disclaimer.//// +//// //// +//// THIS SOFTWARE IS PROVIDED ``AS IS'' AND WITHOUT ANY //// +//// EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED //// +//// TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS //// +//// FOR A PARTICULAR PURPOSE. IN NO EVENT SHALL THE AUTHOR //// +//// OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, //// +//// INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES //// +//// (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE //// +//// GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR //// +//// BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF //// +//// LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT //// +//// (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT //// +//// OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE //// +//// POSSIBILITY OF SUCH DAMAGE. //// +//// //// +///////////////////////////////////////////////////////////////////// + +`timescale 1 ns / 1 ns + +module div_signed(clock, clken, numer, denom, quotient, remain); + + // + // parameters + // + parameter stages = 32; // controls # of pipeline stages + localparam width = stages; + parameter z_width = 2*width; // internally used width + + parameter width_n = width; // width of numerator and quotient <= width + parameter width_d = width; // width of denominator and remainder <= width + + // + // inputs & outputs + // + input clock; // system clock + input clken; // clock enable + + input [width_n -1:0] numer; // dividend + input [width_d -1:0] denom; // divisor + output [width_n -1:0] quotient; // quotient + output [width_d -1:0] remain; // remainder + reg [width_n-1:0] quotient = {(width_n){1'b0}}; + reg [width_d-1:0] remain = {(width_d){1'b0}}; + + // + // functions + // + function [z_width:0] gen_s; + input [z_width:0] si; + input [z_width:0] di; + begin + if(si[z_width]) + gen_s = {si[z_width-1:0], 1'b0} + di; + else + gen_s = {si[z_width-1:0], 1'b0} - di; + end + endfunction + + function [width-1:0] gen_q; + input [width-1:0] qi; + input [z_width:0] si; + begin + gen_q = {qi[width-2:0], ~si[z_width]}; + end + endfunction + + function [width-1:0] assign_s; + input [z_width:0] si; + input [z_width:0] di; + reg [z_width:0] tmp; + begin + if(si[z_width]) + tmp = si + di; + else + tmp = si; + + assign_s = tmp[z_width-1:z_width-width]; + end + endfunction + + // + // variables + // + reg [width-1:0] q_pipe [width-1:0]; + reg [z_width:0] s_pipe [width:0]; + reg [z_width:0] d_pipe [width:0]; + + // added by janders; April 2015 + // spipe1 and spipe2 keep track of whether we need to change the sign + // of the quotient and remainder + // inumer, idenom are the numerator and denominator in absolute value + reg [width-1:0] spipe1; + reg [width-1:0] spipe2; + reg [width-1:0] inumer; + reg [width-1:0] idenom; + + integer n0, n1, n2; + + // This always block takes the absolute value of the numerator + always @(numer) + if (numer[width_n-1]) + inumer = ~numer + 1'b1; + else + inumer = numer; + + // Take the absolute value of the denominator + always @(denom) + if (denom[width_d-1]) + idenom = ~denom + 1'b1; + else + idenom = denom; + + + // generate divisor (d) pipe + always @(idenom) + begin + d_pipe[0] <= {1'b0, idenom, {(z_width-width){1'b0}} }; + end + + always @(posedge clock) + if(clken) + for(n0=1; n0 <= width; n0=n0+1) + d_pipe[n0] <= d_pipe[n0-1]; + + // generate internal remainder pipe + always @(inumer) + begin + s_pipe[0] <= {{(width){1'b0}}, inumer}; + end + + always @(posedge clock) + if(clken) + for(n1=1; n1 <= width; n1=n1+1) + s_pipe[n1] <= gen_s(s_pipe[n1-1], d_pipe[n1-1]); + + // generate quotient pipe + always @(posedge clock) + q_pipe[0] <= 0; + + always @(posedge clock) + if(clken) + for(n2=1; n2 < width; n2=n2+1) + q_pipe[n2] <= gen_q(q_pipe[n2-1], s_pipe[n2]); + + // This is a little tricky. The spipe1 stores whether the + // remainder's should be made negative. In the LLVM srem instruction + // the remainder takes the sign of the numerator. + + // The spipe2 stores whether the quotient is to be made negative. + // This follows the usual math semantics. It should be negative + // if the signs of the dividend and divisor are different (XOR). + + // Added by janders; April 2015 + integer n,m; + always @(posedge clock) + if(clken) + begin + spipe1[0] <= numer[width-1]; + + for(n=1; n < width; n=n+1) + spipe1[n] <= spipe1[n-1]; + + spipe2[0] <= numer[width-1] ^ denom[width-1]; + for(m=1; m < width; m=m+1) + spipe2[m] <= spipe2[m-1]; + end + + always @(posedge clock) + if(clken) // we may need to change the sign of the remainder + remain <= spipe1[width-1] ? ~assign_s(s_pipe[width], d_pipe[width]) + 1'b1 : assign_s(s_pipe[width], d_pipe[width]); + + always @(posedge clock) + if(clken) // we may need to change the sign of the quotient + quotient <= spipe2[width-1] ? ~gen_q(q_pipe[width-1], s_pipe[width]) + 1'b1 : gen_q(q_pipe[width-1], s_pipe[width]); + +endmodule diff --git a/ip/div_signed.v b/ip/div_signed.v deleted file mode 100644 index 7aab57c..0000000 --- a/ip/div_signed.v +++ /dev/null @@ -1,193 +0,0 @@ -///////////////////////////////////////////////////////////////////// -//// //// -//// Non-restoring unsigned divider //// -//// //// -//// Author: Richard Herveille //// -//// richard@asics.ws //// -//// www.asics.ws //// -//// //// -///////////////////////////////////////////////////////////////////// -//// //// -//// Copyright (C) 2002 Richard Herveille //// -//// richard@asics.ws //// -//// //// -//// This source file may be used and distributed without //// -//// restriction provided that this copyright statement is not //// -//// removed from the file and that any derivative work contains //// -//// the original copyright notice and the associated disclaimer.//// -//// //// -//// THIS SOFTWARE IS PROVIDED ``AS IS'' AND WITHOUT ANY //// -//// EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED //// -//// TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS //// -//// FOR A PARTICULAR PURPOSE. IN NO EVENT SHALL THE AUTHOR //// -//// OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, //// -//// INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES //// -//// (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE //// -//// GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR //// -//// BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF //// -//// LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT //// -//// (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT //// -//// OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE //// -//// POSSIBILITY OF SUCH DAMAGE. //// -//// //// -///////////////////////////////////////////////////////////////////// - -`timescale 1 ns / 1 ns - -module div_signed(clock, clken, numer, denom, quotient, remain); - - // - // parameters - // - parameter stages = 32; // controls # of pipeline stages - localparam width = stages; - parameter z_width = 2*width; // internally used width - - parameter width_n = width; // width of numerator and quotient <= width - parameter width_d = width; // width of denominator and remainder <= width - - // - // inputs & outputs - // - input clock; // system clock - input clken; // clock enable - - input [width_n -1:0] numer; // dividend - input [width_d -1:0] denom; // divisor - output [width_n -1:0] quotient; // quotient - output [width_d -1:0] remain; // remainder - reg [width_n-1:0] quotient = {(width_n){1'b0}}; - reg [width_d-1:0] remain = {(width_d){1'b0}}; - - // - // functions - // - function [z_width:0] gen_s; - input [z_width:0] si; - input [z_width:0] di; - begin - if(si[z_width]) - gen_s = {si[z_width-1:0], 1'b0} + di; - else - gen_s = {si[z_width-1:0], 1'b0} - di; - end - endfunction - - function [width-1:0] gen_q; - input [width-1:0] qi; - input [z_width:0] si; - begin - gen_q = {qi[width-2:0], ~si[z_width]}; - end - endfunction - - function [width-1:0] assign_s; - input [z_width:0] si; - input [z_width:0] di; - reg [z_width:0] tmp; - begin - if(si[z_width]) - tmp = si + di; - else - tmp = si; - - assign_s = tmp[z_width-1:z_width-width]; - end - endfunction - - // - // variables - // - reg [width-1:0] q_pipe [width-1:0]; - reg [z_width:0] s_pipe [width:0]; - reg [z_width:0] d_pipe [width:0]; - - // added by janders; April 2015 - // spipe1 and spipe2 keep track of whether we need to change the sign - // of the quotient and remainder - // inumer, idenom are the numerator and denominator in absolute value - reg [width-1:0] spipe1; - reg [width-1:0] spipe2; - reg [width-1:0] inumer; - reg [width-1:0] idenom; - - integer n0, n1, n2; - - // This always block takes the absolute value of the numerator - always @(numer) - if (numer[width_n-1]) - inumer = ~numer + 1'b1; - else - inumer = numer; - - // Take the absolute value of the denominator - always @(denom) - if (denom[width_d-1]) - idenom = ~denom + 1'b1; - else - idenom = denom; - - - // generate divisor (d) pipe - always @(idenom) - begin - d_pipe[0] <= {1'b0, idenom, {(z_width-width){1'b0}} }; - end - - always @(posedge clock) - if(clken) - for(n0=1; n0 <= width; n0=n0+1) - d_pipe[n0] <= d_pipe[n0-1]; - - // generate internal remainder pipe - always @(inumer) - begin - s_pipe[0] <= {{(width){1'b0}}, inumer}; - end - - always @(posedge clock) - if(clken) - for(n1=1; n1 <= width; n1=n1+1) - s_pipe[n1] <= gen_s(s_pipe[n1-1], d_pipe[n1-1]); - - // generate quotient pipe - always @(posedge clock) - q_pipe[0] <= 0; - - always @(posedge clock) - if(clken) - for(n2=1; n2 < width; n2=n2+1) - q_pipe[n2] <= gen_q(q_pipe[n2-1], s_pipe[n2]); - - // This is a little tricky. The spipe1 stores whether the - // remainder's should be made negative. In the LLVM srem instruction - // the remainder takes the sign of the numerator. - - // The spipe2 stores whether the quotient is to be made negative. - // This follows the usual math semantics. It should be negative - // if the signs of the dividend and divisor are different (XOR). - - // Added by janders; April 2015 - integer n,m; - always @(posedge clock) - if(clken) - begin - spipe1[0] <= numer[width-1]; - - for(n=1; n < width; n=n+1) - spipe1[n] <= spipe1[n-1]; - - spipe2[0] <= numer[width-1] ^ denom[width-1]; - for(m=1; m < width; m=m+1) - spipe2[m] <= spipe2[m-1]; - end - - always @(posedge clock) - if(clken) // we may need to change the sign of the remainder - remain <= spipe1[width-1] ? ~assign_s(s_pipe[width], d_pipe[width]) + 1'b1 : assign_s(s_pipe[width], d_pipe[width]); - - always @(posedge clock) - if(clken) // we may need to change the sign of the quotient - quotient <= spipe2[width-1] ? ~gen_q(q_pipe[width-1], s_pipe[width]) + 1'b1 : gen_q(q_pipe[width-1], s_pipe[width]); - -endmodule diff --git a/ip/div_unsigned.sv b/ip/div_unsigned.sv new file mode 100644 index 0000000..db70758 --- /dev/null +++ b/ip/div_unsigned.sv @@ -0,0 +1,146 @@ +///////////////////////////////////////////////////////////////////// +//// //// +//// Non-restoring unsigned divider //// +//// //// +//// Author: Richard Herveille //// +//// richard@asics.ws //// +//// www.asics.ws //// +//// //// +///////////////////////////////////////////////////////////////////// +//// //// +//// Copyright (C) 2002 Richard Herveille //// +//// richard@asics.ws //// +//// //// +//// This source file may be used and distributed without //// +//// restriction provided that this copyright statement is not //// +//// removed from the file and that any derivative work contains //// +//// the original copyright notice and the associated disclaimer.//// +//// //// +//// THIS SOFTWARE IS PROVIDED ``AS IS'' AND WITHOUT ANY //// +//// EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED //// +//// TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS //// +//// FOR A PARTICULAR PURPOSE. IN NO EVENT SHALL THE AUTHOR //// +//// OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, //// +//// INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES //// +//// (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE //// +//// GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR //// +//// BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF //// +//// LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT //// +//// (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT //// +//// OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE //// +//// POSSIBILITY OF SUCH DAMAGE. //// +//// //// +///////////////////////////////////////////////////////////////////// + +`timescale 1 ns / 1 ns + +module div_unsigned(clock, clken, numer, denom, quotient, remain); + + // + // parameters + // + parameter stages = 32; // controls # of pipeline stages + localparam width = stages; + parameter z_width = 2*width; // internally used width + + parameter width_n = width; // width of numerator and quotient <= width + parameter width_d = width; // width of denominator and remainder <= width + + // + // inputs & outputs + // + input clock; // system clock + input clken; // clock enable + + input [width_n -1:0] numer; // dividend + input [width_d -1:0] denom; // divisor + output [width_n -1:0] quotient; // quotient + output [width_d -1:0] remain; // remainder + reg [width_n-1:0] quotient = {(width_n){1'b0}}; + reg [width_d-1:0] remain = {(width_d){1'b0}}; + + // + // functions + // + function [z_width:0] gen_s; + input [z_width:0] si; + input [z_width:0] di; + begin + if(si[z_width]) + gen_s = {si[z_width-1:0], 1'b0} + di; + else + gen_s = {si[z_width-1:0], 1'b0} - di; + end + endfunction + + function [width-1:0] gen_q; + input [width-1:0] qi; + input [z_width:0] si; + begin + gen_q = {qi[width-2:0], ~si[z_width]}; + end + endfunction + + function [width-1:0] assign_s; + input [z_width:0] si; + input [z_width:0] di; + reg [z_width:0] tmp; + begin + if(si[z_width]) + tmp = si + di; + else + tmp = si; + + assign_s = tmp[z_width-1:z_width-width]; + end + endfunction + + // + // variables + // + reg [width-1:0] q_pipe [width-1:0]; + reg [z_width:0] s_pipe [width:0]; + reg [z_width:0] d_pipe [width:0]; + + integer n0, n1, n2; + + // generate divisor (d) pipe + always @(denom) + begin + d_pipe[0] <= {1'b0, {(width-width_d){1'b0}}, denom, {(z_width-width){1'b0}} }; + end + + always @(posedge clock) + if(clken) + for(n0=1; n0 <= width; n0=n0+1) + d_pipe[n0] <= d_pipe[n0-1]; + + // generate internal remainder pipe + always @(numer) + begin + s_pipe[0] <= {{(width){1'b0}}, {(width-width_d){1'b0}}, numer}; + end + + always @(posedge clock) + if(clken) + for(n1=1; n1 <= width; n1=n1+1) + s_pipe[n1] <= gen_s(s_pipe[n1-1], d_pipe[n1-1]); + + // generate quotient pipe + always @(posedge clock) + q_pipe[0] <= 0; + + always @(posedge clock) + if(clken) + for(n2=1; n2 < width; n2=n2+1) + q_pipe[n2] <= gen_q(q_pipe[n2-1], s_pipe[n2]); + + + always @(posedge clock) + if(clken) + quotient <= gen_q(q_pipe[width-1], s_pipe[width]); + + always @(posedge clock) + if(clken) + remain <= assign_s(s_pipe[width], d_pipe[width]); +endmodule diff --git a/ip/div_unsigned.v b/ip/div_unsigned.v deleted file mode 100644 index db70758..0000000 --- a/ip/div_unsigned.v +++ /dev/null @@ -1,146 +0,0 @@ -///////////////////////////////////////////////////////////////////// -//// //// -//// Non-restoring unsigned divider //// -//// //// -//// Author: Richard Herveille //// -//// richard@asics.ws //// -//// www.asics.ws //// -//// //// -///////////////////////////////////////////////////////////////////// -//// //// -//// Copyright (C) 2002 Richard Herveille //// -//// richard@asics.ws //// -//// //// -//// This source file may be used and distributed without //// -//// restriction provided that this copyright statement is not //// -//// removed from the file and that any derivative work contains //// -//// the original copyright notice and the associated disclaimer.//// -//// //// -//// THIS SOFTWARE IS PROVIDED ``AS IS'' AND WITHOUT ANY //// -//// EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED //// -//// TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS //// -//// FOR A PARTICULAR PURPOSE. IN NO EVENT SHALL THE AUTHOR //// -//// OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, //// -//// INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES //// -//// (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE //// -//// GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR //// -//// BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF //// -//// LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT //// -//// (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT //// -//// OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE //// -//// POSSIBILITY OF SUCH DAMAGE. //// -//// //// -///////////////////////////////////////////////////////////////////// - -`timescale 1 ns / 1 ns - -module div_unsigned(clock, clken, numer, denom, quotient, remain); - - // - // parameters - // - parameter stages = 32; // controls # of pipeline stages - localparam width = stages; - parameter z_width = 2*width; // internally used width - - parameter width_n = width; // width of numerator and quotient <= width - parameter width_d = width; // width of denominator and remainder <= width - - // - // inputs & outputs - // - input clock; // system clock - input clken; // clock enable - - input [width_n -1:0] numer; // dividend - input [width_d -1:0] denom; // divisor - output [width_n -1:0] quotient; // quotient - output [width_d -1:0] remain; // remainder - reg [width_n-1:0] quotient = {(width_n){1'b0}}; - reg [width_d-1:0] remain = {(width_d){1'b0}}; - - // - // functions - // - function [z_width:0] gen_s; - input [z_width:0] si; - input [z_width:0] di; - begin - if(si[z_width]) - gen_s = {si[z_width-1:0], 1'b0} + di; - else - gen_s = {si[z_width-1:0], 1'b0} - di; - end - endfunction - - function [width-1:0] gen_q; - input [width-1:0] qi; - input [z_width:0] si; - begin - gen_q = {qi[width-2:0], ~si[z_width]}; - end - endfunction - - function [width-1:0] assign_s; - input [z_width:0] si; - input [z_width:0] di; - reg [z_width:0] tmp; - begin - if(si[z_width]) - tmp = si + di; - else - tmp = si; - - assign_s = tmp[z_width-1:z_width-width]; - end - endfunction - - // - // variables - // - reg [width-1:0] q_pipe [width-1:0]; - reg [z_width:0] s_pipe [width:0]; - reg [z_width:0] d_pipe [width:0]; - - integer n0, n1, n2; - - // generate divisor (d) pipe - always @(denom) - begin - d_pipe[0] <= {1'b0, {(width-width_d){1'b0}}, denom, {(z_width-width){1'b0}} }; - end - - always @(posedge clock) - if(clken) - for(n0=1; n0 <= width; n0=n0+1) - d_pipe[n0] <= d_pipe[n0-1]; - - // generate internal remainder pipe - always @(numer) - begin - s_pipe[0] <= {{(width){1'b0}}, {(width-width_d){1'b0}}, numer}; - end - - always @(posedge clock) - if(clken) - for(n1=1; n1 <= width; n1=n1+1) - s_pipe[n1] <= gen_s(s_pipe[n1-1], d_pipe[n1-1]); - - // generate quotient pipe - always @(posedge clock) - q_pipe[0] <= 0; - - always @(posedge clock) - if(clken) - for(n2=1; n2 < width; n2=n2+1) - q_pipe[n2] <= gen_q(q_pipe[n2-1], s_pipe[n2]); - - - always @(posedge clock) - if(clken) - quotient <= gen_q(q_pipe[width-1], s_pipe[width]); - - always @(posedge clock) - if(clken) - remain <= assign_s(s_pipe[width], d_pipe[width]); -endmodule -- 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 --- Makefile | 2 +- default.nix | 12 ++------ lib/CompCert | 2 +- src/common/Maps.v | 3 +- src/common/Show.v | 30 +++++++------------ src/common/ZExtra.v | 18 ----------- src/hls/AssocMap.v | 86 ++++++++++++++++++++++++++++++----------------------- 7 files changed, 66 insertions(+), 87 deletions(-) diff --git a/Makefile b/Makefile index 9c40941..218598a 100644 --- a/Makefile +++ b/Makefile @@ -6,7 +6,7 @@ ifeq ($(UNAME_S),Darwin) ARCH := verilog-macosx endif -COMPCERTRECDIRS := lib common verilog backend cfrontend driver exportclight cparser +COMPCERTRECDIRS := lib common verilog backend cfrontend driver cparser COQINCLUDES := -R src/common vericert.common \ -R src/extraction vericert.extraction \ diff --git a/default.nix b/default.nix index eabd3d3..e3ce09b 100644 --- a/default.nix +++ b/default.nix @@ -1,7 +1,7 @@ -with import (fetchTarball "https://github.com/NixOS/nixpkgs/archive/1a56d76d718afb6c47dd96602c915b6d23f7c45d.tar.gz") {}; +with import (fetchTarball "https://github.com/NixOS/nixpkgs/archive/00197eff36bb8f7dd7f53a59f730e1fd8e11b1f4.tar.gz") {}; let - ncoq = coq_8_13; - ncoqPackages = coqPackages_8_13; + ncoq = coq_8_14; + ncoqPackages = coqPackages_8_14; in stdenv.mkDerivation { name = "vericert"; @@ -11,12 +11,6 @@ stdenv.mkDerivation { ncoq.ocaml ncoq.ocamlPackages.findlib ncoq.ocamlPackages.menhir ncoq.ocamlPackages.ocamlgraph ncoq.ocamlPackages.merlin ncoq.ocamlPackages.menhirLib - - ncoqPackages.serapi - python3 python3Packages.docutils python3Packages.pygments - python3Packages.dominate - python3Packages.pelican - python3Packages.alectryon ]; enableParallelBuilding = true; diff --git a/lib/CompCert b/lib/CompCert index 1daf96c..4f46759 160000 --- a/lib/CompCert +++ b/lib/CompCert @@ -1 +1 @@ -Subproject commit 1daf96cdca4d828c333cea5c9a314ef861342984 +Subproject commit 4f467596f8674f5f4fbf84a793cb8fcfc35a44a2 diff --git a/src/common/Maps.v b/src/common/Maps.v index f0f264d..13ca276 100644 --- a/src/common/Maps.v +++ b/src/common/Maps.v @@ -17,7 +17,7 @@ * along with this program. If not, see . *) -Set Implicit Arguments. +(*Set Implicit Arguments. Require Export compcert.lib.Maps. @@ -62,3 +62,4 @@ Definition traverse (A B : Type) (f : positive -> A -> res B) m := xtraverse f m Definition traverse1 (A B : Type) (f : A -> res B) := traverse (fun _ => f). End PTree. +*) diff --git a/src/common/Show.v b/src/common/Show.v index 4c66725..20c07e7 100644 --- a/src/common/Show.v +++ b/src/common/Show.v @@ -1,4 +1,4 @@ -(* +(* * Vericert: Verified high-level synthesis. * Copyright (C) 2019-2020 Yann Herklotz * @@ -24,15 +24,11 @@ From Coq Require Import Local Open Scope string. -Class Show A : Type := - { - show : A -> string - }. +Class Show A : Type := { show : A -> string }. +#[export] Instance showBool : Show bool := - { - show := fun b:bool => if b then "true" else "false" - }. + { show := fun b:bool => if b then "true" else "false" }. Fixpoint string_of_nat_aux (time n : nat) (acc : string) : string := let d := match n mod 10 with @@ -52,17 +48,11 @@ Fixpoint string_of_nat_aux (time n : nat) (acc : string) : string := Definition string_of_nat (n : nat) : string := string_of_nat_aux n n "". -Instance showNat : Show nat := - { - show := string_of_nat - }. +#[export] +Instance showNat : Show nat := { show := string_of_nat }. -Instance showZ : Show Z := - { - show a := show (Z.to_nat a) - }. +#[export] +Instance showZ : Show Z := { show a := show (Z.to_nat a) }. -Instance showPositive : Show positive := - { - show a := show (Zpos a) - }. +#[export] +Instance showPositive : Show positive := { show a := show (Zpos a) }. diff --git a/src/common/ZExtra.v b/src/common/ZExtra.v index 62ca54d..3a27cc6 100644 --- a/src/common/ZExtra.v +++ b/src/common/ZExtra.v @@ -45,12 +45,6 @@ Proof. - apply Z.mod_divide; assumption. Qed. -Lemma mod_0_r: forall (m: Z), - m mod 0 = 0. -Proof. - intros. destruct m; reflexivity. -Qed. - Lemma sub_mod_0: forall (a b m: Z), a mod m = 0 -> b mod m = 0 -> @@ -89,18 +83,6 @@ Proof. reflexivity. Qed. -Lemma mod_pow2_same_cases: forall a n, - a mod 2 ^ n = a -> - 2 ^ n = 0 /\ a = 0 \/ 0 <= a < 2 ^ n. -Proof. - intros. - assert (n < 0 \/ 0 <= n) as C by lia. destruct C as [C | C]. - - left. rewrite (Z.pow_neg_r 2 n C) in *. rewrite mod_0_r in H. auto. - - right. - rewrite <- H. apply Z.mod_pos_bound. - apply Z.pow_pos_nonneg; lia. -Qed. - Lemma mod_pow2_same_bounds: forall a n, a mod 2 ^ n = a -> 0 <= n -> 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(-) 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(-) 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 120edd54abf5155b9fc0f3396ee767a6995169e7 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 5 Mar 2022 14:42:11 +0000 Subject: Remove small customisations from license text --- LICENSE | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/LICENSE b/LICENSE index b386211..f288702 100644 --- a/LICENSE +++ b/LICENSE @@ -631,8 +631,8 @@ to attach them to the start of each source file to most effectively state the exclusion of warranty; and each file should have at least the "copyright" line and a pointer to where the full notice is found. - Vericert: Verified high-level synthesis. - Copyright (C) 2019-2021 Yann Herklotz + + Copyright (C) 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 @@ -652,7 +652,7 @@ Also add information on how to contact you by electronic and paper mail. If the program does terminal interaction, make it output a short notice like this when it starts in an interactive mode: - Vericert Copyright (C) 2019-2021 Yann Herklotz + Copyright (C) This program comes with ABSOLUTELY NO WARRANTY; for details type `show w'. This is free software, and you are welcome to redistribute it under certain conditions; type `show c' for details. -- cgit From 518acb94265ad29f7a6081efb682d10bc685e759 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 5 Mar 2022 14:51:00 +0000 Subject: Update links in README --- README.md | 12 ++++++------ README.org | 4 ++-- 2 files changed, 8 insertions(+), 8 deletions(-) diff --git a/README.md b/README.md index e801149..b45a5ab 100644 --- a/README.md +++ b/README.md @@ -4,8 +4,8 @@ - [Setting up Nix](#setting-up-nix) - [Makefile build](#makefile-build) - [Running](#running) - - [Citation](#orge392203) - - [License](#org8a5a598) + - [Citation](#org28fad40) + - [License](#org5eb199f) @@ -109,7 +109,7 @@ To test out `vericert` you can try the following examples which are in the test ``` - + # Citation @@ -132,15 +132,15 @@ If you use Vericert in any way, please cite it using our [OOPSLA’21 paper] ``` - + # License -This project is licensed under [GPLv3](https://www.gnu.org/licenses/gpl-3.0.en.html). The license can be seen in [/LICENSE](file:///LICENSE). +This project is licensed under [GPLv3](https://www.gnu.org/licenses/gpl-3.0.en.html). The license can be seen in [LICENSE](LICENSE). The following external code and its license is present in this repository: -- **[/src/pipelining](file:///src/pipelining):** MIT +- **[src/pipelining](src/pipelining):** MIT ```text Copyright (c) 2008,2009,2010 Jean-Baptiste Tristan and INRIA diff --git a/README.org b/README.org index da29bb1..1569bfb 100644 --- a/README.org +++ b/README.org @@ -133,11 +133,11 @@ If you use Vericert in any way, please cite it using our [[https://yannherklotz. ** License -This project is licensed under [[https://www.gnu.org/licenses/gpl-3.0.en.html][GPLv3]]. The license can be seen in [[/LICENSE][/LICENSE]]. +This project is licensed under [[https://www.gnu.org/licenses/gpl-3.0.en.html][GPLv3]]. The license can be seen in [[file:LICENSE][LICENSE]]. The following external code and its license is present in this repository: -- [[/src/pipelining][/src/pipelining]] :: MIT +- [[file:src/pipelining][src/pipelining]] :: MIT #+begin_src text Copyright (c) 2008,2009,2010 Jean-Baptiste Tristan and INRIA -- 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(+) 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 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 fb9790be528b0b609742b667fca10ee3814a38a6 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 12:17:40 +0000 Subject: Add synthesis processing script --- scripts/synthesise | 98 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 98 insertions(+) create mode 100755 scripts/synthesise diff --git a/scripts/synthesise b/scripts/synthesise new file mode 100755 index 0000000..c865577 --- /dev/null +++ b/scripts/synthesise @@ -0,0 +1,98 @@ +#! /usr/bin/chicken-csi -ss +;; -*- mode: scheme -*- + +(import (chicken port) + (chicken process-context) + args + csv-abnf + matchable + regex + srfi-193 + ssax) + +(define options) +(define operands) + +(define (check-opt b) (cdr (or (assoc b options) `(,b . #f)))) + +(define opts + (list (args:make-option (v verbose) (optional: "LEVEL") "Debug level [default: 0]" + (set! arg (or arg "0"))) + (args:make-option (k keys) (required: "KEY,KEY,...") "Keys to display [default: slice,ramfifo,delay]") + (args:make-option (o output) (required: "FILE") "Output file") + (args:make-option (s suppress) (required: "TYPE,TYPE,...") "Values to suppress from output [default: none]") + (args:make-option (x xml) #:none "Output raw XML from the synthesis tool") + (args:make-option (c csv) #:none "Output processed CSV") + (args:make-option (V version) #:none "Display version" + (print "synthesise v0.1.0") + (exit)) + (args:make-option (h help) #:none "Display this text" + (usage)))) + +(define description + "synthesise: sends a verilog file to be synthesised and returns results as a CSV file.") + +(define (usage) + (with-output-to-port (current-error-port) + (lambda () + (print description) + (newline) + (print "Usage: " (program-name) " [options...] [files...]") + (newline) + (print (args:usage opts)) + (print "Report bugs to git at yannherklotz dot com."))) + (exit 1)) + +(define-values (fmt-cell fmt-record fmt-csv) (make-format ",")) + +(define (map-names n) + (match n + ["XILINX_LUT_FLIP_FLOP_PAIRS_USED" "lut_flip_flop"] + ["XILINX_SLICE" "slice"] + ["XILINX_SLICE_REGISTERS" "regs"] + ["XILINX_SLICE_LUTS" "luts"] + ["XILINX_BLOCK_RAMFIFO" "ramfifo"] + ["XILINX_IOPIN" "iopin"] + ["XILINX_DSPS" "dsps"] + ["XILINX_POWER" "power"] + ["XILINX_DESIGN_DELAY" "delay"] + [_ n])) + +(define (xml-matcher xml) + (match xml + [('*TOP* _ ('document ('application ('section _ . r)))) + (map (match-lambda + [('item ('@ ('value v) ('stringID s))) (list (map-names s) (string->number v))]) r)])) + +(define (parse-xml name file) + (with-input-from-file file + (lambda () + (list name (xml-matcher (ssax:xml->sxml (current-input-port) '())))))) + +;;(define xml-parser +;; (with-input-from-file "nussinov_report.xml" +;; (lambda () (parse-xml "nussinov" (current-input-port))))) + +(define (to-csv-record b head results) + (let ((res (map (lambda (key) + (cadr (assoc key (cadr results)))) head))) + (list->csv-record (if b res (cons (car results) res))))) + +(define (path-to-name path) + (string-substitute "^.*?([^/]+)_report\\.xml$" "\\1" path)) + +(define (write-file file-name text) + (with-output-to-file file-name (lambda () (display text)))) + +(define (convert-files files) + (map (lambda (f) (parse-xml (path-to-name f) f)) files)) + +(define (main args) + (set!-values (options operands) + (args:parse args opts)) + (let ((head (string-split-fields "," (or (check-opt 'keys) "slice,ramfifo,delay") #:infix)) + (suppress (string-split-fields "," (or (check-opt 'suppress) "none") #:infix))) + (let ((body (map (lambda (f) (to-csv-record (member "name" suppress) head f)) + (convert-files operands))) + (header (list->csv-record (if (member "name" suppress) head (cons "name" head))))) + (display (fmt-csv (if (member "header" suppress) body (cons header body))))))) -- cgit From 79c4d9cf28d3aba8c6d5659fbff68283ec41436e Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 14:56:50 +0000 Subject: Add directory traversal to synthesise script --- scripts/synthesise | 58 +++++++++++++++++++++++++++++++++++++----------------- 1 file changed, 40 insertions(+), 18 deletions(-) diff --git a/scripts/synthesise b/scripts/synthesise index c865577..4e8dbec 100755 --- a/scripts/synthesise +++ b/scripts/synthesise @@ -3,10 +3,11 @@ (import (chicken port) (chicken process-context) + (chicken string) + (chicken irregex) + (chicken file) args - csv-abnf matchable - regex srfi-193 ssax) @@ -29,6 +30,7 @@ (args:make-option (h help) #:none "Display this text" (usage)))) +(: description string) (define description "synthesise: sends a verilog file to be synthesised and returns results as a CSV file.") @@ -43,8 +45,6 @@ (print "Report bugs to git at yannherklotz dot com."))) (exit 1)) -(define-values (fmt-cell fmt-record fmt-csv) (make-format ",")) - (define (map-names n) (match n ["XILINX_LUT_FLIP_FLOP_PAIRS_USED" "lut_flip_flop"] @@ -58,6 +58,13 @@ ["XILINX_DESIGN_DELAY" "delay"] [_ n])) +(define (csv:fmt-row l) (string-intersperse (map ->string l) ",")) + +(define (csv:fmt-table-string l) (apply string-append (map (lambda (s) (string-append s "\n")) l))) + +(define (csv:fmt-table l) (apply string-append (map (lambda (s) (string-append s "\n")) + (map csv:fmt-row l)))) + (define (xml-matcher xml) (match xml [('*TOP* _ ('document ('application ('section _ . r)))) @@ -69,30 +76,45 @@ (lambda () (list name (xml-matcher (ssax:xml->sxml (current-input-port) '())))))) -;;(define xml-parser -;; (with-input-from-file "nussinov_report.xml" -;; (lambda () (parse-xml "nussinov" (current-input-port))))) - (define (to-csv-record b head results) (let ((res (map (lambda (key) (cadr (assoc key (cadr results)))) head))) - (list->csv-record (if b res (cons (car results) res))))) + (csv:fmt-row (if b res (cons (car results) res))))) +(: path-to-name (string -> string)) (define (path-to-name path) - (string-substitute "^.*?([^/]+)_report\\.xml$" "\\1" path)) - -(define (write-file file-name text) - (with-output-to-file file-name (lambda () (display text)))) + (irregex-replace "^.*?([^/]+)_report\\.xml$" path 1)) (define (convert-files files) (map (lambda (f) (parse-xml (path-to-name f) f)) files)) +(: split-at-comma (string -> (list-of string))) +(define (split-at-comma s) (string-split s ",")) + +(: find-all-xml (string -> (list-of string))) +(define (find-all-xml dir) (find-files dir #:test ".*\\.xml$")) + +(define (get-files-from-op operands) + (match operands + [(d) (cond [(directory-exists? d) (find-all-xml d)] + [else (list d)])] + [_ operands])) + +(define (with-output thk) + (if (check-opt 'output) + (with-output-to-file (check-opt 'output) thk) + (thk))) + (define (main args) (set!-values (options operands) (args:parse args opts)) - (let ((head (string-split-fields "," (or (check-opt 'keys) "slice,ramfifo,delay") #:infix)) - (suppress (string-split-fields "," (or (check-opt 'suppress) "none") #:infix))) + (let ((head (split-at-comma (or (check-opt 'keys) "slice,ramfifo,delay"))) + (suppress (split-at-comma (or (check-opt 'suppress) "none"))) + (files (get-files-from-op operands))) (let ((body (map (lambda (f) (to-csv-record (member "name" suppress) head f)) - (convert-files operands))) - (header (list->csv-record (if (member "name" suppress) head (cons "name" head))))) - (display (fmt-csv (if (member "header" suppress) body (cons header body))))))) + (convert-files files))) + (header (csv:fmt-row (if (member "name" suppress) head (cons "name" head))))) + (with-output + (lambda () + (display (csv:fmt-table-string + (if (member "header" suppress) body (cons header body))))))))) -- 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(-) 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 314592d5593faa02a0330451a22fd9b8cc251b8d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 14:57:23 +0000 Subject: Add synthesis bash scripts --- scripts/synth-ssh.sh | 27 +++++++++++++++++++++++++++ scripts/synth-ssh0.sh | 27 +++++++++++++++++++++++++++ scripts/synth.sh | 32 ++++++++++++++++++++++++++++++++ 3 files changed, 86 insertions(+) create mode 100755 scripts/synth-ssh.sh create mode 100755 scripts/synth-ssh0.sh create mode 100755 scripts/synth.sh diff --git a/scripts/synth-ssh.sh b/scripts/synth-ssh.sh new file mode 100755 index 0000000..fffb9f9 --- /dev/null +++ b/scripts/synth-ssh.sh @@ -0,0 +1,27 @@ +#!/usr/bin/bash + +# Assumes that the Verilog is passed on the command line, that the tcl file is in synth.tcl and +# returns encode_report.xml. + +scriptsdir=$(dirname "$(readlink -f "$BASH_SOURCE")") + +bench=$1 +output=$2 +machine=ee-beholder1.ee.ic.ac.uk +user=ymh15 +files="$scriptsdir/synth.tcl $output/$bench.v" +log="$output/${bench}_synth.log" + +date >$log + +temp=$(ssh $user@$machine "mktemp -d" 2>>$log) + +>&2 echo "synthesising $bench $temp" +rsync $files $user@$machine:$temp/ >>$log 2>&1 +ssh $user@$machine \ + "bash -lc 'cd $temp && cp $(basename $bench).v main.v && vivado -mode batch -source synth.tcl'" \ + >>$log 2>&1 +rsync $user@$machine:$temp/encode_report.xml $output/${bench}_report.xml >>$log 2>&1 +ssh $user@$machine "rm -rf '$temp'" >>$log 2>&1 +rm -f main.v >>$log 2>&1 +>&2 echo "done $bench" diff --git a/scripts/synth-ssh0.sh b/scripts/synth-ssh0.sh new file mode 100755 index 0000000..977309c --- /dev/null +++ b/scripts/synth-ssh0.sh @@ -0,0 +1,27 @@ +#!/usr/bin/bash + +# Assumes that the Verilog is passed on the command line, that the tcl file is in synth.tcl and +# returns encode_report.xml. + +scriptsdir=$(dirname "$(readlink -f "$BASH_SOURCE")") + +bench=$1 +output=$2 +machine=ee-beholder0.ee.ic.ac.uk +user=ymh15 +files="$scriptsdir/synth.tcl $output/$bench.v" +log="$output/${bench}_synth.log" + +date >$log + +temp=$(ssh $user@$machine "mktemp -d" 2>>$log) + +>&2 echo "synthesising $bench $temp" +rsync $files $user@$machine:$temp/ >>$log 2>&1 +ssh $user@$machine \ + "bash -lc 'cd $temp && cp $(basename $bench).v main.v && vivado -mode batch -source synth.tcl'" \ + >>$log 2>&1 +rsync $user@$machine:$temp/encode_report.xml $output/${bench}_report.xml >>$log 2>&1 +ssh $user@$machine "rm -rf '$temp'" >>$log 2>&1 +rm -f main.v >>$log 2>&1 +>&2 echo "done $bench" diff --git a/scripts/synth.sh b/scripts/synth.sh new file mode 100755 index 0000000..79d7164 --- /dev/null +++ b/scripts/synth.sh @@ -0,0 +1,32 @@ +#!/usr/bin/bash + +set -x + +scriptsdir=$(dirname "$(readlink -f "$BASH_SOURCE")") + +if [[ -z "$1" ]]; then + parallel=1 +else + parallel=$1 +fi + +if [[ -z "$2" ]]; then + output=$(pwd) +else + output=$2 +fi + +if [[ -z "$3" ]]; then + source=$(pwd) +else + source=$3 +fi + +echo "copying directory structure from $source to $output" +mkdir -p $output +rsync -am --include '*/' --include '*.v' --exclude '*' $source/ $output/ + +echo "executing $parallel runs in parallel" +cat $scriptsdir/../benchmarks/polybench-syn/benchmark-list-master | \ + xargs --max-procs=$parallel --replace=% \ + $scriptsdir/synth-ssh0.sh % $output -- cgit From adc981f648a5f513b7501715e86dcb1ec7b4f986 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 15:03:28 +0000 Subject: Rename into scheme --- scripts/synthesis-results.scm | 122 ++++++++++++++++++++++++++++++++++++++++++ scripts/synthesise | 120 ----------------------------------------- 2 files changed, 122 insertions(+), 120 deletions(-) create mode 100755 scripts/synthesis-results.scm delete mode 100755 scripts/synthesise diff --git a/scripts/synthesis-results.scm b/scripts/synthesis-results.scm new file mode 100755 index 0000000..faf8919 --- /dev/null +++ b/scripts/synthesis-results.scm @@ -0,0 +1,122 @@ +#! /usr/bin/chicken-csi -s +;; -*- mode: scheme -*- + +(import (chicken port) + (chicken process-context) + (chicken string) + (chicken irregex) + (chicken file) + args + matchable + srfi-193 + ssax) + +(define options) +(define operands) + +(define (check-opt b) (cdr (or (assoc b options) `(,b . #f)))) + +(define opts + (list (args:make-option (v verbose) (optional: "LEVEL") "Debug level [default: 0]" + (set! arg (or arg "0"))) + (args:make-option (k keys) (required: "KEY,KEY,...") "Keys to display [default: slice,ramfifo,delay]") + (args:make-option (o output) (required: "FILE") "Output file") + (args:make-option (s suppress) (required: "TYPE,TYPE,...") "Values to suppress from output [default: none]") + (args:make-option (x xml) #:none "Output raw XML from the synthesis tool") + (args:make-option (c csv) #:none "Output processed CSV") + (args:make-option (V version) #:none "Display version" + (print "synthesise v0.1.0") + (exit)) + (args:make-option (h help) #:none "Display this text" + (usage)))) + +(: description string) +(define description + "synthesise: sends a verilog file to be synthesised and returns results as a CSV file.") + +(define (usage) + (with-output-to-port (current-error-port) + (lambda () + (print description) + (newline) + (print "Usage: " (program-name) " [options...] [files...]") + (newline) + (print (args:usage opts)) + (print "Report bugs to git at yannherklotz dot com."))) + (exit 1)) + +(define (map-names n) + (match n + ["XILINX_LUT_FLIP_FLOP_PAIRS_USED" "lut_flip_flop"] + ["XILINX_SLICE" "slice"] + ["XILINX_SLICE_REGISTERS" "regs"] + ["XILINX_SLICE_LUTS" "luts"] + ["XILINX_BLOCK_RAMFIFO" "ramfifo"] + ["XILINX_IOPIN" "iopin"] + ["XILINX_DSPS" "dsps"] + ["XILINX_POWER" "power"] + ["XILINX_DESIGN_DELAY" "delay"] + [_ n])) + +(define (csv:fmt-row l) (string-intersperse (map ->string l) ",")) + +(define (csv:fmt-table-string l) (apply string-append (map (lambda (s) (string-append s "\n")) l))) + +(define (csv:fmt-table l) (apply string-append (map (lambda (s) (string-append s "\n")) + (map csv:fmt-row l)))) + +(define (xml-matcher xml) + (match xml + [('*TOP* _ ('document ('application ('section _ . r)))) + (map (match-lambda + [('item ('@ ('value v) ('stringID s))) (list (map-names s) (string->number v))]) r)])) + +(define (parse-xml name file) + (with-input-from-file file + (lambda () + (list name (xml-matcher (ssax:xml->sxml (current-input-port) '())))))) + +(define (to-csv-record b head results) + (let ((res (map (lambda (key) + (cadr (assoc key (cadr results)))) head))) + (csv:fmt-row (if b res (cons (car results) res))))) + +(: path-to-name (string -> string)) +(define (path-to-name path) + (irregex-replace "^.*?([^/]+)_report\\.xml$" path 1)) + +(define (convert-files files) + (map (lambda (f) (parse-xml (path-to-name f) f)) files)) + +(: split-at-comma (string -> (list-of string))) +(define (split-at-comma s) (string-split s ",")) + +(: find-all-xml (string -> (list-of string))) +(define (find-all-xml dir) (find-files dir #:test ".*\\.xml$")) + +(define (get-files-from-op operands) + (match operands + [(d) (cond [(directory-exists? d) (find-all-xml d)] + [else (list d)])] + [_ operands])) + +(define (with-output thk) + (if (check-opt 'output) + (with-output-to-file (check-opt 'output) thk) + (thk))) + +(define (main args) + (set!-values (options operands) + (args:parse args opts)) + (let ((head (split-at-comma (or (check-opt 'keys) "slice,ramfifo,delay"))) + (suppress (split-at-comma (or (check-opt 'suppress) "none"))) + (files (get-files-from-op operands))) + (let ((body (map (lambda (f) (to-csv-record (member "name" suppress) head f)) + (convert-files files))) + (header (csv:fmt-row (if (member "name" suppress) head (cons "name" head))))) + (with-output + (lambda () + (display (csv:fmt-table-string + (if (member "header" suppress) body (cons header body))))))))) + +(main (command-line-arguments)) diff --git a/scripts/synthesise b/scripts/synthesise deleted file mode 100755 index 4e8dbec..0000000 --- a/scripts/synthesise +++ /dev/null @@ -1,120 +0,0 @@ -#! /usr/bin/chicken-csi -ss -;; -*- mode: scheme -*- - -(import (chicken port) - (chicken process-context) - (chicken string) - (chicken irregex) - (chicken file) - args - matchable - srfi-193 - ssax) - -(define options) -(define operands) - -(define (check-opt b) (cdr (or (assoc b options) `(,b . #f)))) - -(define opts - (list (args:make-option (v verbose) (optional: "LEVEL") "Debug level [default: 0]" - (set! arg (or arg "0"))) - (args:make-option (k keys) (required: "KEY,KEY,...") "Keys to display [default: slice,ramfifo,delay]") - (args:make-option (o output) (required: "FILE") "Output file") - (args:make-option (s suppress) (required: "TYPE,TYPE,...") "Values to suppress from output [default: none]") - (args:make-option (x xml) #:none "Output raw XML from the synthesis tool") - (args:make-option (c csv) #:none "Output processed CSV") - (args:make-option (V version) #:none "Display version" - (print "synthesise v0.1.0") - (exit)) - (args:make-option (h help) #:none "Display this text" - (usage)))) - -(: description string) -(define description - "synthesise: sends a verilog file to be synthesised and returns results as a CSV file.") - -(define (usage) - (with-output-to-port (current-error-port) - (lambda () - (print description) - (newline) - (print "Usage: " (program-name) " [options...] [files...]") - (newline) - (print (args:usage opts)) - (print "Report bugs to git at yannherklotz dot com."))) - (exit 1)) - -(define (map-names n) - (match n - ["XILINX_LUT_FLIP_FLOP_PAIRS_USED" "lut_flip_flop"] - ["XILINX_SLICE" "slice"] - ["XILINX_SLICE_REGISTERS" "regs"] - ["XILINX_SLICE_LUTS" "luts"] - ["XILINX_BLOCK_RAMFIFO" "ramfifo"] - ["XILINX_IOPIN" "iopin"] - ["XILINX_DSPS" "dsps"] - ["XILINX_POWER" "power"] - ["XILINX_DESIGN_DELAY" "delay"] - [_ n])) - -(define (csv:fmt-row l) (string-intersperse (map ->string l) ",")) - -(define (csv:fmt-table-string l) (apply string-append (map (lambda (s) (string-append s "\n")) l))) - -(define (csv:fmt-table l) (apply string-append (map (lambda (s) (string-append s "\n")) - (map csv:fmt-row l)))) - -(define (xml-matcher xml) - (match xml - [('*TOP* _ ('document ('application ('section _ . r)))) - (map (match-lambda - [('item ('@ ('value v) ('stringID s))) (list (map-names s) (string->number v))]) r)])) - -(define (parse-xml name file) - (with-input-from-file file - (lambda () - (list name (xml-matcher (ssax:xml->sxml (current-input-port) '())))))) - -(define (to-csv-record b head results) - (let ((res (map (lambda (key) - (cadr (assoc key (cadr results)))) head))) - (csv:fmt-row (if b res (cons (car results) res))))) - -(: path-to-name (string -> string)) -(define (path-to-name path) - (irregex-replace "^.*?([^/]+)_report\\.xml$" path 1)) - -(define (convert-files files) - (map (lambda (f) (parse-xml (path-to-name f) f)) files)) - -(: split-at-comma (string -> (list-of string))) -(define (split-at-comma s) (string-split s ",")) - -(: find-all-xml (string -> (list-of string))) -(define (find-all-xml dir) (find-files dir #:test ".*\\.xml$")) - -(define (get-files-from-op operands) - (match operands - [(d) (cond [(directory-exists? d) (find-all-xml d)] - [else (list d)])] - [_ operands])) - -(define (with-output thk) - (if (check-opt 'output) - (with-output-to-file (check-opt 'output) thk) - (thk))) - -(define (main args) - (set!-values (options operands) - (args:parse args opts)) - (let ((head (split-at-comma (or (check-opt 'keys) "slice,ramfifo,delay"))) - (suppress (split-at-comma (or (check-opt 'suppress) "none"))) - (files (get-files-from-op operands))) - (let ((body (map (lambda (f) (to-csv-record (member "name" suppress) head f)) - (convert-files files))) - (header (csv:fmt-row (if (member "name" suppress) head (cons "name" head))))) - (with-output - (lambda () - (display (csv:fmt-table-string - (if (member "header" suppress) body (cons header body))))))))) -- cgit From 921c8ace00ad57440467d495bdb19407a377193f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 15:14:53 +0000 Subject: Add Makefile in scripts directory --- scripts/Makefile | 9 +++++++++ 1 file changed, 9 insertions(+) create mode 100644 scripts/Makefile diff --git a/scripts/Makefile b/scripts/Makefile new file mode 100644 index 0000000..bf51ce7 --- /dev/null +++ b/scripts/Makefile @@ -0,0 +1,9 @@ +all: synthesis + +synthesis: synthesis-results.scm + chicken-csc -static -o $@ $< + +install: synthesis + install -s $< ~/.local/bin + +.PHONY: all install -- cgit From a1c958a50129342aeccf38e0e591bed6c394cf74 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 15:16:28 +0000 Subject: Remove more unnecessary scripts --- scripts/convert.sh | 14 ------ scripts/download_artifact.sh | 8 ---- scripts/gather_data.rkt | 111 ------------------------------------------- 3 files changed, 133 deletions(-) delete mode 100755 scripts/convert.sh delete mode 100755 scripts/download_artifact.sh delete mode 100644 scripts/gather_data.rkt diff --git a/scripts/convert.sh b/scripts/convert.sh deleted file mode 100755 index c2ef311..0000000 --- a/scripts/convert.sh +++ /dev/null @@ -1,14 +0,0 @@ -file=$1 -benchmark=$(echo $1 | sed 's:.*/\([^/]\+\)/encode_report.xml:\1:') - -lut_flip_flop=$(sed -n "s/.*XILINX_LUT_FLIP_FLOP_PAIRS_USED.*\"\([0-9.]*\)\".*/\1/p" $file) -slice=$(sed -n "s/.*XILINX_SLICE\".*\"\([0-9.]*\)\".*/\1/p" $file) -regs=$(sed -n "s/.*XILINX_SLICE_REGISTERS.*\"\([0-9.]*\)\".*/\1/p" $file) -luts=$(sed -n "s/.*XILINX_SLICE_LUTS.*\"\([0-9.]*\)\".*/\1/p" $file) -ramfifo=$(sed -n "s/.*XILINX_BLOCK_RAMFIFO.*\"\([0-9.]*\)\".*/\1/p" $file) -iopin=$(sed -n "s/.*XILINX_IOPIN.*\"\([0-9.]*\)\".*/\1/p" $file) -dsps=$(sed -n "s/.*XILINX_DSPS.*\"\([0-9.]*\)\".*/\1/p" $file) -power=$(sed -n "s/.*XILINX_POWER.*\"\([0-9.]*\)\".*/\1/p" $file) -delay=$(sed -n "s/.*XILINX_DESIGN_DELAY.*\"\([0-9.]*\)\".*/\1/p" $file) - -echo $benchmark,$lut_flip_flop,$slice,$regs,$luts,$ramfifo,$iopin,$dsps,$power,$delay >>synth.csv diff --git a/scripts/download_artifact.sh b/scripts/download_artifact.sh deleted file mode 100755 index cac69f6..0000000 --- a/scripts/download_artifact.sh +++ /dev/null @@ -1,8 +0,0 @@ -#!/bin/sh - -mkdir -p docs/html/proof -cd docs/html/proof -curl -v -L -H "Accept: application/vnd.github.v3+json" -H "Authorization: token $GITHUB_TOKEN" https://api.github.com/repos/ymherklotz/vericert/actions/artifacts/26286264/zip -o html-documentation.zip -unzip html-documentation.zip -rm html-documentation.zip -cp ../../css/coqdoc.css . diff --git a/scripts/gather_data.rkt b/scripts/gather_data.rkt deleted file mode 100644 index c2da7d4..0000000 --- a/scripts/gather_data.rkt +++ /dev/null @@ -1,111 +0,0 @@ -#lang racket - -(require racket/match) -(require racket/list) -(require racket/file) - -(require threading) - -(require xml) -(require xml/path) - -(require csv-reading) -(require csv-writing) - -(require graphite) -(require data-frame) -(require sawzall) - -(permissive-xexprs #t) - -(define (to-name f) (regexp-replace #rx".*/([^/]+)" f "\\1")) - -(define (write-file file data) - (with-output-to-file file - (lambda () - (display data)) - #:exists 'replace)) - -(define files (file->lines "../benchmarks/polybench-syn/benchmark-list-master")) -(define base "/home/ymherklotz/projects/mpardalos-vericert/results-vericert-fun-full-inlining") - -(define (list->hash l) - (foldr (lambda (v l) - (hash-set l (car v) (cadr v))) - (hash) l)) - -(define (list->hashn l) - (foldr (lambda (v l) - (hash-set l (car v) (string->number (cadr v)))) - (hash) l)) - -(define name-f-map - (list->hash (map (lambda (f) (list (to-name f) (string-append base "/" f "_report.xml"))) files))) - -(define (parse-vivado-report f) - (let* ([encode-xml-port (open-input-file f)] - [report (xml->xexpr (document-element (read-xml encode-xml-port)))]) - (close-input-port encode-xml-port) - report)) - -(define (process-vivado-report report) - (let ([maps (map (lambda (x) (match x [(list e (list (list a b) (list c d))) (list b d)])) - (filter-not string? (se-path*/list '(section) report)))]) - (list->hashn maps))) - -(define (vivado-report-f f) (process-vivado-report (parse-vivado-report f))) - -;;(vivado-report-f "./data/data-mining/covariance/encode_report.xml") -;; -;;(vivado-report-f "/home/ymherklotz/projects/mpardalos-vericert/results-vericert-fun-full-inlining/medley/nussinov_report.xml") - -(define synth-f (map flatten (hash-map name-f-map - (lambda (n f) (let ([x (vivado-report-f f)]) - (list n - (hash-ref x "XILINX_LUT_FLIP_FLOP_PAIRS_USED") - (hash-ref x "XILINX_SLICE") - (hash-ref x "XILINX_SLICE_REGISTERS") - (hash-ref x "XILINX_SLICE_LUTS") - (hash-ref x "XILINX_BLOCK_RAMFIFO") - (hash-ref x "XILINX_IOPIN") - (hash-ref x "XILINX_DSPS") - (hash-ref x "XILINX_POWER") - (hash-ref x "XILINX_DESIGN_DELAY"))))))) - -(define (parse-sim-report f) - (let* ([exec-csv (open-input-file f)] - [report (csv->list exec-csv)]) - (close-input-port exec-csv) - report)) - -(define sim-report (list->hashn (parse-sim-report (string-append base "/exec.csv")))) - -(write-file "out.csv" - (table->string - (append '((benchmark lut_flip_flop slice regs luts ramfifo iopin dsps power delay cycles)) - (map (lambda (x) (append x (list (hash-ref sim-report (car x))))) synth-f)))) - -;;(define vivado-report (process-vivado-report (parse-vivado-report "encode_report.xml"))) -;;(define cycles (hash-ref sim-report "covariance")) -;;(define delay (hash-ref vivado-report "XILINX_DESIGN_DELAY")) -;; -;;(define gss (df-read/csv "exec.csv")) -;; -;;(df-add-series gss (make-series "tool" #:data (make-vector (df-row-count gss) "vericert"))) -;; -;;(df-add-series gss (make-series "delay" #:data (list->vector (map (lambda (x) -;; (~> "/encode_report.xml" -;; (string-append x _) -;; parse-vivado-report -;; process-vivado-report -;; (hash-ref "XILINX_DESIGN_DELAY"))) -;; (filter-not (lambda (x) (equal? x "test-case")) (map car (hash->list sim-report))))))) -;; -;;(show gss everything #:n-rows 'all) -;; -;;(graph #:data gss -;; #:title "Chart" -;; #:mapping (aes #:x "test-case" #:y "cycle count") -;; (col #:gap 0.25)) -;; -;;;(define csv (open-output-file "out.csv")) -- cgit From a7a8e4481bce810d077ea682b8379535a20931c9 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 21:09:26 +0000 Subject: Update documentation files --- Makefile | 11 ++++++++--- docs/Makefile | 14 +++++++++++--- docs/res/publish-html.el | 22 ---------------------- docs/res/publish-man.el | 12 ------------ docs/res/publish-manual.el | 12 ------------ docs/res/publish.el | 23 ++++++++++++++++++++++ scripts/Makefile | 15 ++++++++++++--- scripts/synthesis-results.org | 44 +++++++++++++++++++++++++++++++++++++++++++ scripts/synthesis-results.scm | 2 +- 9 files changed, 99 insertions(+), 56 deletions(-) delete mode 100644 docs/res/publish-html.el delete mode 100644 docs/res/publish-man.el delete mode 100644 docs/res/publish-manual.el create mode 100644 docs/res/publish.el create mode 100644 scripts/synthesis-results.org diff --git a/Makefile b/Makefile index 218598a..03bfaa5 100644 --- a/Makefile +++ b/Makefile @@ -38,11 +38,13 @@ lib/COMPCERTSTAMP: lib/CompCert/Makefile.config $(MAKE) HAS_RUNTIME_LIB=false CLIGHTGEN=false INSTALL_COQDEV=false -C lib/CompCert touch $@ -install: - install -d $(PREFIX)/bin +install: docs/vericert.1 sed -i'' -e 's/arch=verilog/arch=x86/' _build/default/driver/compcert.ini - install -C _build/default/driver/compcert.ini $(PREFIX)/bin/. + install -d $(PREFIX)/bin + install -C _build/default/driver/compcert.ini $(PREFIX)/bin install -C _build/default/driver/VericertDriver.exe $(PREFIX)/bin/vericert + install -d $(PREFIX)/share/man/man1 + install -C $< $(PREFIX)/share/man/man1 proof: Makefile.coq $(MAKE) -f Makefile.coq @@ -75,6 +77,9 @@ Makefile.coq: @echo "COQMAKE Makefile.coq" $(COQBIN)coq_makefile $(COQINCLUDES) $(VS) -o Makefile.coq +docs/vericert.1: + $(MAKE) -C docs vericert.1 + clean:: Makefile.coq $(MAKE) -f Makefile.coq clean $(MAKE) -C test clean diff --git a/docs/Makefile b/docs/Makefile index 93083cd..9d4f361 100644 --- a/docs/Makefile +++ b/docs/Makefile @@ -3,19 +3,27 @@ all: manual src man-html install-deps: emacs --batch --load ./res/install-deps.el +%.man: %.org + emacs --batch --file $< --load ./res/publish.el --funcall org-man-export-to-man + +%.html: %.org + emacs --batch --file $< --load ./res/publish.el --funcall org-html-export-to-html + manual: mkdir -p manual - emacs --batch --file documentation.org --load ./res/publish-manual.el + emacs --batch --file documentation.org --load ./res/publish.el --funcall org-texinfo-export-to-texinfo makeinfo --html --number-sections --no-split \ --css-ref "https://www.gnu.org/software/emacs/manual.css" \ vericert.texi -o ./manual/index.html cp -r images ./manual/. -man-html: +man-html: man.html mkdir -p man - emacs --batch --file man.org --load ./res/publish-html.el cp man.html ./man/vericert.1.html +vericert.1: man.man + cp $< $@ + src: $(MAKE) -C .. doc cp -r ../html src diff --git a/docs/res/publish-html.el b/docs/res/publish-html.el deleted file mode 100644 index a9c7a9d..0000000 --- a/docs/res/publish-html.el +++ /dev/null @@ -1,22 +0,0 @@ -(require 'package) -(package-initialize) - -(require 'org) -(require 'org-transclusion) -(require 'ox) -(require 'ox-html) -(require 'htmlize) - -(setq org-transclusion-exclude-elements nil - org-html-head-include-default-style nil - org-html-head-include-scripts nil - org-html-postamble-format '(("en" "")) - org-html-postamble t - org-html-divs '((preamble "header" "header") - (content "article" "content") - (postamble "footer" "postamble")) - org-html-doctype "html5" - org-html-htmlize-output-type 'css) - -(org-transclusion-add-all) -(org-html-export-to-html) diff --git a/docs/res/publish-man.el b/docs/res/publish-man.el deleted file mode 100644 index 634b454..0000000 --- a/docs/res/publish-man.el +++ /dev/null @@ -1,12 +0,0 @@ -(require 'package) -(package-initialize) - -(require 'org) -(require 'org-transclusion) -(require 'ox) -(require 'ox-man) - -(setq org-transclusion-exclude-elements nil) - -(org-transclusion-add-all) -(org-man-export-to-man) diff --git a/docs/res/publish-manual.el b/docs/res/publish-manual.el deleted file mode 100644 index df5db6b..0000000 --- a/docs/res/publish-manual.el +++ /dev/null @@ -1,12 +0,0 @@ -(require 'package) -(package-initialize) - -(require 'org) -(require 'org-transclusion) -(require 'ox) -(require 'ox-texinfo) - -(setq org-transclusion-exclude-elements nil) - -(org-transclusion-add-all) -(org-texinfo-export-to-texinfo) diff --git a/docs/res/publish.el b/docs/res/publish.el new file mode 100644 index 0000000..c083eb0 --- /dev/null +++ b/docs/res/publish.el @@ -0,0 +1,23 @@ +(require 'package) +(package-initialize) + +(require 'org) +(require 'org-transclusion) +(require 'ox) +(require 'ox-html) +(require 'htmlize) +(require 'ox-texinfo) +(require 'ox-man) + +(setq org-transclusion-exclude-elements nil + org-html-head-include-default-style nil + org-html-head-include-scripts nil + org-html-postamble-format '(("en" "")) + org-html-postamble t + org-html-divs '((preamble "header" "header") + (content "article" "content") + (postamble "footer" "postamble")) + org-html-doctype "html5" + org-html-htmlize-output-type 'css) + +(org-transclusion-add-all) diff --git a/scripts/Makefile b/scripts/Makefile index bf51ce7..e31a6fd 100644 --- a/scripts/Makefile +++ b/scripts/Makefile @@ -1,9 +1,18 @@ +PREFIX ?= .. + all: synthesis -synthesis: synthesis-results.scm +%: %.scm chicken-csc -static -o $@ $< -install: synthesis - install -s $< ~/.local/bin +%.1: %.org + emacs --batch --file $< --load ../docs/res/publish.el --funcall org-man-export-to-man + cp $(<:.org=.man) $@ + +install: synthesis-results synthesis-results.1 + install -d $(PREFIX)/bin + install -C synthesis-results $(PREFIX)/bin + install -d $(PREFIX)/share/man/man1 + install -C synthesis-results.1 $(PREFIX)/share/man/man1 .PHONY: all install diff --git a/scripts/synthesis-results.org b/scripts/synthesis-results.org new file mode 100644 index 0000000..602b2ba --- /dev/null +++ b/scripts/synthesis-results.org @@ -0,0 +1,44 @@ +#+title: synthesis-results +#+man_class_options: :section-id "1" +#+options: toc:nil num:nil +#+html_head_extra: + +* NAME + +synthesis-results: sends a verilog file to be synthesised and returns results as a CSV file. + +* SYNOPSYS + +Usage: *synthesis-results* [options...] [files...] + +* DESCRIPTION + +- -v, --verbose[=LEVEL] :: Debug level [default: 0] +- -k, --keys=KEY,KEY,... :: Keys to display [default: slice,ramfifo,delay] +- -o, --output=FILE :: Output file +- -s, --suppress=TYPE,TYPE,... :: Values to suppress from output [default: none] +- -x, --xml :: Output raw XML from the synthesis tool +- -c, --csv :: Output processed CSV +- -V, --version :: Display version +- -h, --help :: Display this text + +* AUTHOR + +Written by Yann Herklotz. + +* COPYRIGHT + +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 . diff --git a/scripts/synthesis-results.scm b/scripts/synthesis-results.scm index faf8919..3a87d58 100755 --- a/scripts/synthesis-results.scm +++ b/scripts/synthesis-results.scm @@ -32,7 +32,7 @@ (: description string) (define description - "synthesise: sends a verilog file to be synthesised and returns results as a CSV file.") + "synthesis-results: sends a verilog file to be synthesised and returns results as a CSV file.") (define (usage) (with-output-to-port (current-error-port) -- cgit From eb95db64b07f8f9daac9944751803506fdfbd577 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 21:24:11 +0000 Subject: Clean up scripts some more --- scripts/Makefile | 6 +++++- scripts/run-legup.sh | 0 scripts/run-vivado.sh | 8 -------- scripts/synth-ssh.sh | 7 ++++--- scripts/synth-ssh0.sh | 27 --------------------------- scripts/synth.sh | 2 +- 6 files changed, 10 insertions(+), 40 deletions(-) mode change 100644 => 100755 scripts/run-legup.sh delete mode 100755 scripts/run-vivado.sh delete mode 100755 scripts/synth-ssh0.sh diff --git a/scripts/Makefile b/scripts/Makefile index e31a6fd..f6421f4 100644 --- a/scripts/Makefile +++ b/scripts/Makefile @@ -7,7 +7,7 @@ all: synthesis %.1: %.org emacs --batch --file $< --load ../docs/res/publish.el --funcall org-man-export-to-man - cp $(<:.org=.man) $@ + mv $(<:.org=.man) $@ install: synthesis-results synthesis-results.1 install -d $(PREFIX)/bin @@ -15,4 +15,8 @@ install: synthesis-results synthesis-results.1 install -d $(PREFIX)/share/man/man1 install -C synthesis-results.1 $(PREFIX)/share/man/man1 +clean: + rm -f synthesis-results synthesis-results.1 + rm -f *.link + .PHONY: all install diff --git a/scripts/run-legup.sh b/scripts/run-legup.sh old mode 100644 new mode 100755 diff --git a/scripts/run-vivado.sh b/scripts/run-vivado.sh deleted file mode 100755 index 117054d..0000000 --- a/scripts/run-vivado.sh +++ /dev/null @@ -1,8 +0,0 @@ -#!/bin/bash - -benchmark=./$1/$2 -echo $benchmark - - cp ./synth.tcl $benchmark/. 2>/dev/null - cd $benchmark || exit 1 - vivado -mode batch -source synth.tcl >vivado.log 2>&1 diff --git a/scripts/synth-ssh.sh b/scripts/synth-ssh.sh index fffb9f9..a6ce349 100755 --- a/scripts/synth-ssh.sh +++ b/scripts/synth-ssh.sh @@ -5,9 +5,10 @@ scriptsdir=$(dirname "$(readlink -f "$BASH_SOURCE")") -bench=$1 -output=$2 -machine=ee-beholder1.ee.ic.ac.uk +num=$1 +bench=$2 +output=$3 +machine=ee-beholder${num}.ee.ic.ac.uk user=ymh15 files="$scriptsdir/synth.tcl $output/$bench.v" log="$output/${bench}_synth.log" diff --git a/scripts/synth-ssh0.sh b/scripts/synth-ssh0.sh deleted file mode 100755 index 977309c..0000000 --- a/scripts/synth-ssh0.sh +++ /dev/null @@ -1,27 +0,0 @@ -#!/usr/bin/bash - -# Assumes that the Verilog is passed on the command line, that the tcl file is in synth.tcl and -# returns encode_report.xml. - -scriptsdir=$(dirname "$(readlink -f "$BASH_SOURCE")") - -bench=$1 -output=$2 -machine=ee-beholder0.ee.ic.ac.uk -user=ymh15 -files="$scriptsdir/synth.tcl $output/$bench.v" -log="$output/${bench}_synth.log" - -date >$log - -temp=$(ssh $user@$machine "mktemp -d" 2>>$log) - ->&2 echo "synthesising $bench $temp" -rsync $files $user@$machine:$temp/ >>$log 2>&1 -ssh $user@$machine \ - "bash -lc 'cd $temp && cp $(basename $bench).v main.v && vivado -mode batch -source synth.tcl'" \ - >>$log 2>&1 -rsync $user@$machine:$temp/encode_report.xml $output/${bench}_report.xml >>$log 2>&1 -ssh $user@$machine "rm -rf '$temp'" >>$log 2>&1 -rm -f main.v >>$log 2>&1 ->&2 echo "done $bench" diff --git a/scripts/synth.sh b/scripts/synth.sh index 79d7164..b1c2696 100755 --- a/scripts/synth.sh +++ b/scripts/synth.sh @@ -29,4 +29,4 @@ rsync -am --include '*/' --include '*.v' --exclude '*' $source/ $output/ echo "executing $parallel runs in parallel" cat $scriptsdir/../benchmarks/polybench-syn/benchmark-list-master | \ xargs --max-procs=$parallel --replace=% \ - $scriptsdir/synth-ssh0.sh % $output + $scriptsdir/synth-ssh.sh 0 % $output -- cgit From 52b050695d489512e2118f1535348f18359322db Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 21:26:29 +0000 Subject: Add a license to synthesis-results --- scripts/synthesis-results.scm | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/scripts/synthesis-results.scm b/scripts/synthesis-results.scm index 3a87d58..11c66c1 100755 --- a/scripts/synthesis-results.scm +++ b/scripts/synthesis-results.scm @@ -1,5 +1,20 @@ #! /usr/bin/chicken-csi -s ;; -*- mode: scheme -*- +;; +;; 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 . (import (chicken port) (chicken process-context) -- cgit From 23fe30f8c23ecb743880cb9239410eb51bf1abab Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 21:37:51 +0000 Subject: Update permissions and .gitignore --- .gitignore | 8 ++++++++ Makefile | 4 ++-- scripts/Makefile | 2 +- 3 files changed, 11 insertions(+), 3 deletions(-) diff --git a/.gitignore b/.gitignore index 43c9acc..4acfbb0 100644 --- a/.gitignore +++ b/.gitignore @@ -82,3 +82,11 @@ creduce_bug_*/ /docs/man/ /docs/manual/ /docs/src/ + +*~ +*.man +*.link + +/bin +/share +/results* diff --git a/Makefile b/Makefile index 03bfaa5..e9df0dd 100644 --- a/Makefile +++ b/Makefile @@ -41,10 +41,10 @@ lib/COMPCERTSTAMP: lib/CompCert/Makefile.config install: docs/vericert.1 sed -i'' -e 's/arch=verilog/arch=x86/' _build/default/driver/compcert.ini install -d $(PREFIX)/bin - install -C _build/default/driver/compcert.ini $(PREFIX)/bin + install -C -m 644 _build/default/driver/compcert.ini $(PREFIX)/bin install -C _build/default/driver/VericertDriver.exe $(PREFIX)/bin/vericert install -d $(PREFIX)/share/man/man1 - install -C $< $(PREFIX)/share/man/man1 + install -C -m 644 $< $(PREFIX)/share/man/man1 proof: Makefile.coq $(MAKE) -f Makefile.coq diff --git a/scripts/Makefile b/scripts/Makefile index f6421f4..0b5ff33 100644 --- a/scripts/Makefile +++ b/scripts/Makefile @@ -13,7 +13,7 @@ install: synthesis-results synthesis-results.1 install -d $(PREFIX)/bin install -C synthesis-results $(PREFIX)/bin install -d $(PREFIX)/share/man/man1 - install -C synthesis-results.1 $(PREFIX)/share/man/man1 + install -C -m 644 synthesis-results.1 $(PREFIX)/share/man/man1 clean: rm -f synthesis-results synthesis-results.1 -- cgit From 10d614dca61aeccc9cb2bc0f2e4c77aae99e7a00 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 --- .gitignore | 4 ++ benchmarks/polybench-syn-div/exec.csv | 27 ------------ benchmarks/polybench-syn-div/poly.csv | 26 ----------- benchmarks/polybench-syn-div/quartus_synth.tcl | 35 --------------- benchmarks/polybench-syn-div/run-vericert.sh | 41 ----------------- benchmarks/polybench-syn-div/script.R | 29 ------------ benchmarks/polybench-syn-div/setup-syn-vericert.sh | 24 ---------- benchmarks/polybench-syn-div/syn-remote.sh | 51 ---------------------- benchmarks/polybench-syn/poly.csv | 26 ----------- benchmarks/polybench-syn/quartus_synth.tcl | 35 --------------- benchmarks/polybench-syn/run-vericert.sh | 46 ------------------- benchmarks/polybench-syn/script.R | 29 ------------ benchmarks/polybench-syn/setup-syn-vericert.sh | 24 ---------- benchmarks/polybench-syn/syn-remote.sh | 51 ---------------------- scripts/quartus_synth.tcl | 35 +++++++++++++++ scripts/run-vericert.sh | 46 +++++++++++++++++++ scripts/script.R | 29 ++++++++++++ scripts/syn-remote.sh | 51 ++++++++++++++++++++++ scripts/synthesis-results.scm | 5 +-- 19 files changed, 167 insertions(+), 447 deletions(-) delete mode 100644 benchmarks/polybench-syn-div/exec.csv delete mode 100644 benchmarks/polybench-syn-div/poly.csv delete mode 100644 benchmarks/polybench-syn-div/quartus_synth.tcl delete mode 100755 benchmarks/polybench-syn-div/run-vericert.sh delete mode 100644 benchmarks/polybench-syn-div/script.R delete mode 100755 benchmarks/polybench-syn-div/setup-syn-vericert.sh delete mode 100755 benchmarks/polybench-syn-div/syn-remote.sh delete mode 100644 benchmarks/polybench-syn/poly.csv delete mode 100644 benchmarks/polybench-syn/quartus_synth.tcl delete mode 100755 benchmarks/polybench-syn/run-vericert.sh delete mode 100644 benchmarks/polybench-syn/script.R delete mode 100755 benchmarks/polybench-syn/setup-syn-vericert.sh delete mode 100755 benchmarks/polybench-syn/syn-remote.sh create mode 100644 scripts/quartus_synth.tcl create mode 100755 scripts/run-vericert.sh create mode 100644 scripts/script.R create mode 100755 scripts/syn-remote.sh diff --git a/.gitignore b/.gitignore index 4acfbb0..586f3a7 100644 --- a/.gitignore +++ b/.gitignore @@ -1,3 +1,7 @@ +/benchmarks/**/* +!/benchmarks/**/*.* +!/benchmarks/**/*/ + .*.aux *.a *.cma diff --git a/benchmarks/polybench-syn-div/exec.csv b/benchmarks/polybench-syn-div/exec.csv deleted file mode 100644 index e28109b..0000000 --- a/benchmarks/polybench-syn-div/exec.csv +++ /dev/null @@ -1,27 +0,0 @@ -adi,1422354 -heat-3d,580770 -fdtd-2d,901430 -jacobi-1d,19622 -seidel-2d,664780 -jacobi-2d,344072 -nussinov,954402 -floyd-warshall,5373798 -3mm,536114 -2mm,404478 -doitgen,351988 -bicg,53916 -mvt,70204 -atas,58424 -syrk,271816 -gemver,117394 -symm,240172 -gesummv,37700 -gemm,328104 -trmm,144688 -syr2k,436520 -cholesky,2535686 -trisolv,25192 -lu,2853646 -ludcmp,2601382 -durbin,22974 -covariance,288392 diff --git a/benchmarks/polybench-syn-div/poly.csv b/benchmarks/polybench-syn-div/poly.csv deleted file mode 100644 index fe71983..0000000 --- a/benchmarks/polybench-syn-div/poly.csv +++ /dev/null @@ -1,26 +0,0 @@ -benchmark,legupcycles,legupfreqMHz,leguplogicutilisation,legupregs,leguprams,legupdsps,legupcomptime,vericertcycles,vericertfreqMHz,vericertlogicutilisation,vericertregs,vericertrams,vericertdsps,vericertcomptime -durbin,15106,188.61,2509,4008,0,8,4.77,19852,199.2,3027,6168,0,8,0.077 -lu,482766,244.62,3116,4593,0,10,4.72,2634766,92.97,54806,106037,0,6,0.097 -ludcmp,470843,249.69,3605,5397,0,15,4.87,2401354,83.52,57145,111408,0,10,0.132 -trisolv,35382,213.9,2412,3749,0,3,4.73,33550,112.59,28101,55109,0,2,0.086 -2mm,60088,226.5,1114,1936,0,7,4.80,427098,116.9,32600,63212,0,18,0.136 -3mm,204195,188.96,4210,4801,0,43,4.88,539430,97.13,45073,89719,0,18,0.147 -atas,126288,193.24,3026,3719,0,10,4.80,92000,130.41,28603,56646,0,6,0.069 -bicg,11907,303.4,308,537,0,6,4.78,121134,122.74,30091,58799,0,8,0.097 -mvt,16806,384.47,372,597,0,4,4.79,139194,119.93,30753,60450,0,6,0.130 -doitgen,57199,252.14,909,1402,0,2,5.05,350302,126.18,19898,38461,0,4,0.105 -symm,64903,284.41,2155,3170,0,10,4.63,248930,113.92,27693,54514,0,14,0.115 -syrk,57395,278.01,598,976,0,2,4.73,309018,87.15,76889,57816,0,8,0.094 -syr2k,125705,240.85,3149,3679,0,6,4.85,478040,78.6,120116,82032,0,12,0.116 -trmm,41432,281.61,610,990,0,4,4.71,147528,105.61,64031,40502,0,4,0.089 -gemm,83676,192.68,1029,1544,0,35,4.79,360772,121.61,31853,62126,0,16,0.104 -gemver,28087,303.49,1854,2380,0,8,4.68,175326,107.27,32615,64118,0,14,0.099 -gesummv,6634,310.46,298,504,0,4,4.77,111094,79.97,113876,72908,0,10,0.101 -covariance,109992,245.16,2098,3096,0,5,4.77,297450,110.57,28729,56660,0,4,0.083 -fdtd-2d,214153,262.61,2736,3801,0,2,4.73,831912,108.23,31333,61421,0,6,0.116 -heat-3d,41059,115.54,3132,2910,0,60,4.92,555930,110.42,33915,67273,0,9,0.181 -jacobi-1d,6914,386.25,1355,1885,0,0,4.72,16606,277.93,1636,3305,0,0,0.071 -jacobi-2d,84609,240.79,2347,3185,0,2,4.81,357100,113.53,30393,59782,0,4,0.079 -seidel-2d,345294,232.4,2128,3337,0,2,4.68,875758,127.99,26948,53133,0,2,0.091 -floyd-warshall,1238764,235.52,1869,2367,0,2,4.71,4762766,109.4,59859,118101,0,2,0.094 -nussinov,216467,273.07,1078,1431,0,2,4.79,837958,90.73,60663,119303,0,0,0.080 diff --git a/benchmarks/polybench-syn-div/quartus_synth.tcl b/benchmarks/polybench-syn-div/quartus_synth.tcl deleted file mode 100644 index 6edbf0c..0000000 --- a/benchmarks/polybench-syn-div/quartus_synth.tcl +++ /dev/null @@ -1,35 +0,0 @@ -# PRiME pre-KAPow kernel flow -# Performs pre-KAPow run steps for instrumenting arbitrary Verilog for power monitoring -# James Davis, 2015 - -load_package flow - -project_new -overwrite syn -set_global_assignment -name FAMILY "Arria 10" -set_global_assignment -name DEVICE 10AX115H4F34E3LG -set_global_assignment -name SYSTEMVERILOG_FILE top.v -set_global_assignment -name TOP_LEVEL_ENTITY main -#set_global_assignment -name SDC_FILE syn.sdc -#set_global_assignment -name auto_resource_sharing on -#set_global_assignment -name enable_state_machine_inference on -#set_global_assignment -name optimization_technique area -#set_global_assignment -name synthesis_effort fast -#set_global_assignment -name AUTO_RAM_RECOGNITION on -#set_global_assignment -name remove_duplicate_registers on -#set_instance_assignment -name RAMSTYLE_ATTRIBUTE LOGIC -to ram - -execute_module -tool map - -execute_module -tool fit - -execute_module -tool sta - -#execute_module -tool eda -args "--simulation --tool=vcs" - -# set_global_assignment -name POWER_OUTPUT_SAF_NAME ${kernel}.asf -# set_global_assignment -name POWER_DEFAULT_INPUT_IO_TOGGLE_RATE "12.5 %" -# set_global_assignment -name POWER_REPORT_SIGNAL_ACTIVITY ON -# set_global_assignment -name POWER_REPORT_POWER_DISSIPATION ON -# execute_module -tool pow - -project_close diff --git a/benchmarks/polybench-syn-div/run-vericert.sh b/benchmarks/polybench-syn-div/run-vericert.sh deleted file mode 100755 index 6cf4cd9..0000000 --- a/benchmarks/polybench-syn-div/run-vericert.sh +++ /dev/null @@ -1,41 +0,0 @@ -#!/usr/bin/env bash - -rm exec.csv - -top=$(pwd) - #set up -while read benchmark ; do - echo "Running "$benchmark - ./$benchmark.gcc > $benchmark.clog - cresult=$(cat $benchmark.clog | cut -d' ' -f2) - echo "C output: "$cresult - ./$benchmark.iver > $benchmark.tmp - veriresult=$(tail -1 $benchmark.tmp | cut -d' ' -f2) - cycles=$(tail -2 $benchmark.tmp | head -1 | tr -s ' ' | cut -d' ' -f2) - echo "Verilog output: "$veriresult - - #Undefined checks - if test -z $veriresult - then - echo "FAIL: Verilog returned nothing" - #exit 0 - fi - - # Don't care checks - if [ $veriresult == "x" ] - then - echo "FAIL: Verilog returned don't cares" - #exit 0 - fi - - # unequal result check - if [ $cresult -ne $veriresult ] - then - echo "FAIL: Verilog and C output do not match!" - #exit 0 - else - echo "PASS" - fi - name=$(echo $benchmark | awk -v FS="/" '{print $NF}') - echo $name","$cycles >> exec.csv -done < benchmark-list-master diff --git a/benchmarks/polybench-syn-div/script.R b/benchmarks/polybench-syn-div/script.R deleted file mode 100644 index 0be16da..0000000 --- a/benchmarks/polybench-syn-div/script.R +++ /dev/null @@ -1,29 +0,0 @@ -library("psych") - -data = read.csv("poly.csv", header=TRUE) -leguptime = (data$legupcycles/data$legupfreqMHz) -veritime = data$vericertcycles/data$vericertfreqMHz -print(lm(veritime ~ leguptime)) -leguputil = data$leguplogicutilisation/427200*100 -veriutil = data$vericertlogicutilisation/427200*100 -print(lm (veriutil ~ leguputil)) -legupct = data$legupcomptime -verict = data$vericertcomptime -print(lm ( verict ~ legupct )) - -cycleslowdown=data$vericertcycles/data$legupcycles - -print("Cycle count slow down") -print(geometric.mean(cycleslowdown)) -print("Wall clock slow down") -print(geometric.mean(veritime/leguptime)) -print("Area overhead") -print(geometric.mean(veriutil/leguputil)) -print("Compilation time speedup") -print(geometric.mean(legupct/verict)) -print("LegUp RAM use") -print(geometric.mean(data$legupregs)) -print("Vericert RAM use") -print(geometric.mean(data$vericertregs)) -print("Area overhead") -print(geometric.mean(data$vericertregs/data$legupregs)) diff --git a/benchmarks/polybench-syn-div/setup-syn-vericert.sh b/benchmarks/polybench-syn-div/setup-syn-vericert.sh deleted file mode 100755 index 22356f7..0000000 --- a/benchmarks/polybench-syn-div/setup-syn-vericert.sh +++ /dev/null @@ -1,24 +0,0 @@ -#! /bin/bash - -top=$(pwd) - #set up - basedir=poly-syn - sshhost=$1 - ssh $sshhost "cd ~; rm -r $basedir" - ssh $sshhost "cd ~; mkdir $basedir" - scp quartus_synth.tcl $sshhost:$basedir - scp syn-remote.sh $sshhost:$basedir - rm syn-list - - while read benchmark ; - do - echo "Copying "$benchmark" over" - name=$(echo $benchmark | awk -v FS="/" '{print $NF}') - echo "Name: "$name - benchdir="~/$basedir/$name" - scp $benchmark.v $sshhost:~/$basedir - echo $name >> syn-list - done < benchmark-list-master - - # copy list over - scp syn-list $sshhost:$basedir diff --git a/benchmarks/polybench-syn-div/syn-remote.sh b/benchmarks/polybench-syn-div/syn-remote.sh deleted file mode 100755 index 879db2e..0000000 --- a/benchmarks/polybench-syn-div/syn-remote.sh +++ /dev/null @@ -1,51 +0,0 @@ -#! /bin/bash - -#setup -while read benchmark ; -do -echo "Setting up "$benchmark -rm -r $benchmark -mkdir $benchmark -cp $benchmark.v $benchmark/top.v - -done < syn-list - -#synthesis - -count=0 -while read benchmark ; - -do -echo "Synthesising "$benchmark -cd $benchmark -quartus_sh -t ../quartus_synth.tcl & -let "count=count+1" -cd .. - -if [ $count -eq 4 ] -then -echo "I am here" -wait -count=0 -fi - -done < syn-list - -if [ $count -lt 4 ] -then -wait -fi - -#extract -while read benchmark ; do - cd $benchmark - echo $(pwd) - freq=$(grep MHz syn.sta.rpt | tail -2 | head -1 | awk '{print $2}') - lut=$(sed -n -e 8p syn.fit.summary | awk '{print $6}' | sed 's/,//g') - regs=$(sed -n -e 9p syn.fit.summary | awk '{print $4}') - bram=$(sed -n -e 13p syn.fit.summary | awk '{print $5}') - dsp=$(sed -n -e 14p syn.fit.summary | awk '{print $5}') - cd .. - echo $benchmark","$freq","$lut","$regs","$bram","$dsp >> results -done < syn-list - diff --git a/benchmarks/polybench-syn/poly.csv b/benchmarks/polybench-syn/poly.csv deleted file mode 100644 index fe71983..0000000 --- a/benchmarks/polybench-syn/poly.csv +++ /dev/null @@ -1,26 +0,0 @@ -benchmark,legupcycles,legupfreqMHz,leguplogicutilisation,legupregs,leguprams,legupdsps,legupcomptime,vericertcycles,vericertfreqMHz,vericertlogicutilisation,vericertregs,vericertrams,vericertdsps,vericertcomptime -durbin,15106,188.61,2509,4008,0,8,4.77,19852,199.2,3027,6168,0,8,0.077 -lu,482766,244.62,3116,4593,0,10,4.72,2634766,92.97,54806,106037,0,6,0.097 -ludcmp,470843,249.69,3605,5397,0,15,4.87,2401354,83.52,57145,111408,0,10,0.132 -trisolv,35382,213.9,2412,3749,0,3,4.73,33550,112.59,28101,55109,0,2,0.086 -2mm,60088,226.5,1114,1936,0,7,4.80,427098,116.9,32600,63212,0,18,0.136 -3mm,204195,188.96,4210,4801,0,43,4.88,539430,97.13,45073,89719,0,18,0.147 -atas,126288,193.24,3026,3719,0,10,4.80,92000,130.41,28603,56646,0,6,0.069 -bicg,11907,303.4,308,537,0,6,4.78,121134,122.74,30091,58799,0,8,0.097 -mvt,16806,384.47,372,597,0,4,4.79,139194,119.93,30753,60450,0,6,0.130 -doitgen,57199,252.14,909,1402,0,2,5.05,350302,126.18,19898,38461,0,4,0.105 -symm,64903,284.41,2155,3170,0,10,4.63,248930,113.92,27693,54514,0,14,0.115 -syrk,57395,278.01,598,976,0,2,4.73,309018,87.15,76889,57816,0,8,0.094 -syr2k,125705,240.85,3149,3679,0,6,4.85,478040,78.6,120116,82032,0,12,0.116 -trmm,41432,281.61,610,990,0,4,4.71,147528,105.61,64031,40502,0,4,0.089 -gemm,83676,192.68,1029,1544,0,35,4.79,360772,121.61,31853,62126,0,16,0.104 -gemver,28087,303.49,1854,2380,0,8,4.68,175326,107.27,32615,64118,0,14,0.099 -gesummv,6634,310.46,298,504,0,4,4.77,111094,79.97,113876,72908,0,10,0.101 -covariance,109992,245.16,2098,3096,0,5,4.77,297450,110.57,28729,56660,0,4,0.083 -fdtd-2d,214153,262.61,2736,3801,0,2,4.73,831912,108.23,31333,61421,0,6,0.116 -heat-3d,41059,115.54,3132,2910,0,60,4.92,555930,110.42,33915,67273,0,9,0.181 -jacobi-1d,6914,386.25,1355,1885,0,0,4.72,16606,277.93,1636,3305,0,0,0.071 -jacobi-2d,84609,240.79,2347,3185,0,2,4.81,357100,113.53,30393,59782,0,4,0.079 -seidel-2d,345294,232.4,2128,3337,0,2,4.68,875758,127.99,26948,53133,0,2,0.091 -floyd-warshall,1238764,235.52,1869,2367,0,2,4.71,4762766,109.4,59859,118101,0,2,0.094 -nussinov,216467,273.07,1078,1431,0,2,4.79,837958,90.73,60663,119303,0,0,0.080 diff --git a/benchmarks/polybench-syn/quartus_synth.tcl b/benchmarks/polybench-syn/quartus_synth.tcl deleted file mode 100644 index 6edbf0c..0000000 --- a/benchmarks/polybench-syn/quartus_synth.tcl +++ /dev/null @@ -1,35 +0,0 @@ -# PRiME pre-KAPow kernel flow -# Performs pre-KAPow run steps for instrumenting arbitrary Verilog for power monitoring -# James Davis, 2015 - -load_package flow - -project_new -overwrite syn -set_global_assignment -name FAMILY "Arria 10" -set_global_assignment -name DEVICE 10AX115H4F34E3LG -set_global_assignment -name SYSTEMVERILOG_FILE top.v -set_global_assignment -name TOP_LEVEL_ENTITY main -#set_global_assignment -name SDC_FILE syn.sdc -#set_global_assignment -name auto_resource_sharing on -#set_global_assignment -name enable_state_machine_inference on -#set_global_assignment -name optimization_technique area -#set_global_assignment -name synthesis_effort fast -#set_global_assignment -name AUTO_RAM_RECOGNITION on -#set_global_assignment -name remove_duplicate_registers on -#set_instance_assignment -name RAMSTYLE_ATTRIBUTE LOGIC -to ram - -execute_module -tool map - -execute_module -tool fit - -execute_module -tool sta - -#execute_module -tool eda -args "--simulation --tool=vcs" - -# set_global_assignment -name POWER_OUTPUT_SAF_NAME ${kernel}.asf -# set_global_assignment -name POWER_DEFAULT_INPUT_IO_TOGGLE_RATE "12.5 %" -# set_global_assignment -name POWER_REPORT_SIGNAL_ACTIVITY ON -# set_global_assignment -name POWER_REPORT_POWER_DISSIPATION ON -# execute_module -tool pow - -project_close diff --git a/benchmarks/polybench-syn/run-vericert.sh b/benchmarks/polybench-syn/run-vericert.sh deleted file mode 100755 index 9deaa10..0000000 --- a/benchmarks/polybench-syn/run-vericert.sh +++ /dev/null @@ -1,46 +0,0 @@ -#!/usr/bin/env bash - -rm exec.csv - -top=$(pwd) -#set up -while read benchmark ; do - printf "%10s\t" $(echo "$benchmark" | sed -e 's|/| |g') - ./$benchmark.gcc > $benchmark.clog - cresult=$(cat $benchmark.clog | cut -d' ' -f2) - #echo "C output: "$cresult - #./$benchmark.iver > $benchmark.tmp - if [[ ! -f ./$benchmark.verilator/Vmain ]]; then - echo -e "\e[0;91mFAIL\e[0m: Verilog failed compilation" - continue - fi - ./$benchmark.verilator/Vmain > $benchmark.tmp - veriresult=$(tail -1 $benchmark.tmp | cut -d' ' -f2) - cycles=$(tail -2 $benchmark.tmp | head -1 | tr -s ' ' | cut -d' ' -f2) - #echo "Verilog output: "$veriresult - - #Undefined checks - if [[ -z "$veriresult" ]] - then - echo "\e[0;91mFAIL\e[0m: Verilog returned nothing" - #exit 0 - fi - - # Don't care checks - if [[ $veriresult == "x" ]] - then - echo "\e[0;91mFAIL\e[0m: Verilog returned don't cares" - #exit 0 - fi - - # unequal result check - if [[ $cresult -ne $veriresult ]] - then - echo -e "\e[0;91mFAIL\e[0m: Verilog and C output do not match!" - #exit 0 - else - echo -e "\e[0;92mPASS\e[0m" - fi - name=$(echo $benchmark | awk -v FS="/" '{print $NF}') - echo $name","$cycles >> exec.csv -done < benchmark-list-master diff --git a/benchmarks/polybench-syn/script.R b/benchmarks/polybench-syn/script.R deleted file mode 100644 index 0be16da..0000000 --- a/benchmarks/polybench-syn/script.R +++ /dev/null @@ -1,29 +0,0 @@ -library("psych") - -data = read.csv("poly.csv", header=TRUE) -leguptime = (data$legupcycles/data$legupfreqMHz) -veritime = data$vericertcycles/data$vericertfreqMHz -print(lm(veritime ~ leguptime)) -leguputil = data$leguplogicutilisation/427200*100 -veriutil = data$vericertlogicutilisation/427200*100 -print(lm (veriutil ~ leguputil)) -legupct = data$legupcomptime -verict = data$vericertcomptime -print(lm ( verict ~ legupct )) - -cycleslowdown=data$vericertcycles/data$legupcycles - -print("Cycle count slow down") -print(geometric.mean(cycleslowdown)) -print("Wall clock slow down") -print(geometric.mean(veritime/leguptime)) -print("Area overhead") -print(geometric.mean(veriutil/leguputil)) -print("Compilation time speedup") -print(geometric.mean(legupct/verict)) -print("LegUp RAM use") -print(geometric.mean(data$legupregs)) -print("Vericert RAM use") -print(geometric.mean(data$vericertregs)) -print("Area overhead") -print(geometric.mean(data$vericertregs/data$legupregs)) diff --git a/benchmarks/polybench-syn/setup-syn-vericert.sh b/benchmarks/polybench-syn/setup-syn-vericert.sh deleted file mode 100755 index 22356f7..0000000 --- a/benchmarks/polybench-syn/setup-syn-vericert.sh +++ /dev/null @@ -1,24 +0,0 @@ -#! /bin/bash - -top=$(pwd) - #set up - basedir=poly-syn - sshhost=$1 - ssh $sshhost "cd ~; rm -r $basedir" - ssh $sshhost "cd ~; mkdir $basedir" - scp quartus_synth.tcl $sshhost:$basedir - scp syn-remote.sh $sshhost:$basedir - rm syn-list - - while read benchmark ; - do - echo "Copying "$benchmark" over" - name=$(echo $benchmark | awk -v FS="/" '{print $NF}') - echo "Name: "$name - benchdir="~/$basedir/$name" - scp $benchmark.v $sshhost:~/$basedir - echo $name >> syn-list - done < benchmark-list-master - - # copy list over - scp syn-list $sshhost:$basedir diff --git a/benchmarks/polybench-syn/syn-remote.sh b/benchmarks/polybench-syn/syn-remote.sh deleted file mode 100755 index 879db2e..0000000 --- a/benchmarks/polybench-syn/syn-remote.sh +++ /dev/null @@ -1,51 +0,0 @@ -#! /bin/bash - -#setup -while read benchmark ; -do -echo "Setting up "$benchmark -rm -r $benchmark -mkdir $benchmark -cp $benchmark.v $benchmark/top.v - -done < syn-list - -#synthesis - -count=0 -while read benchmark ; - -do -echo "Synthesising "$benchmark -cd $benchmark -quartus_sh -t ../quartus_synth.tcl & -let "count=count+1" -cd .. - -if [ $count -eq 4 ] -then -echo "I am here" -wait -count=0 -fi - -done < syn-list - -if [ $count -lt 4 ] -then -wait -fi - -#extract -while read benchmark ; do - cd $benchmark - echo $(pwd) - freq=$(grep MHz syn.sta.rpt | tail -2 | head -1 | awk '{print $2}') - lut=$(sed -n -e 8p syn.fit.summary | awk '{print $6}' | sed 's/,//g') - regs=$(sed -n -e 9p syn.fit.summary | awk '{print $4}') - bram=$(sed -n -e 13p syn.fit.summary | awk '{print $5}') - dsp=$(sed -n -e 14p syn.fit.summary | awk '{print $5}') - cd .. - echo $benchmark","$freq","$lut","$regs","$bram","$dsp >> results -done < syn-list - diff --git a/scripts/quartus_synth.tcl b/scripts/quartus_synth.tcl new file mode 100644 index 0000000..6edbf0c --- /dev/null +++ b/scripts/quartus_synth.tcl @@ -0,0 +1,35 @@ +# PRiME pre-KAPow kernel flow +# Performs pre-KAPow run steps for instrumenting arbitrary Verilog for power monitoring +# James Davis, 2015 + +load_package flow + +project_new -overwrite syn +set_global_assignment -name FAMILY "Arria 10" +set_global_assignment -name DEVICE 10AX115H4F34E3LG +set_global_assignment -name SYSTEMVERILOG_FILE top.v +set_global_assignment -name TOP_LEVEL_ENTITY main +#set_global_assignment -name SDC_FILE syn.sdc +#set_global_assignment -name auto_resource_sharing on +#set_global_assignment -name enable_state_machine_inference on +#set_global_assignment -name optimization_technique area +#set_global_assignment -name synthesis_effort fast +#set_global_assignment -name AUTO_RAM_RECOGNITION on +#set_global_assignment -name remove_duplicate_registers on +#set_instance_assignment -name RAMSTYLE_ATTRIBUTE LOGIC -to ram + +execute_module -tool map + +execute_module -tool fit + +execute_module -tool sta + +#execute_module -tool eda -args "--simulation --tool=vcs" + +# set_global_assignment -name POWER_OUTPUT_SAF_NAME ${kernel}.asf +# set_global_assignment -name POWER_DEFAULT_INPUT_IO_TOGGLE_RATE "12.5 %" +# set_global_assignment -name POWER_REPORT_SIGNAL_ACTIVITY ON +# set_global_assignment -name POWER_REPORT_POWER_DISSIPATION ON +# execute_module -tool pow + +project_close diff --git a/scripts/run-vericert.sh b/scripts/run-vericert.sh new file mode 100755 index 0000000..9deaa10 --- /dev/null +++ b/scripts/run-vericert.sh @@ -0,0 +1,46 @@ +#!/usr/bin/env bash + +rm exec.csv + +top=$(pwd) +#set up +while read benchmark ; do + printf "%10s\t" $(echo "$benchmark" | sed -e 's|/| |g') + ./$benchmark.gcc > $benchmark.clog + cresult=$(cat $benchmark.clog | cut -d' ' -f2) + #echo "C output: "$cresult + #./$benchmark.iver > $benchmark.tmp + if [[ ! -f ./$benchmark.verilator/Vmain ]]; then + echo -e "\e[0;91mFAIL\e[0m: Verilog failed compilation" + continue + fi + ./$benchmark.verilator/Vmain > $benchmark.tmp + veriresult=$(tail -1 $benchmark.tmp | cut -d' ' -f2) + cycles=$(tail -2 $benchmark.tmp | head -1 | tr -s ' ' | cut -d' ' -f2) + #echo "Verilog output: "$veriresult + + #Undefined checks + if [[ -z "$veriresult" ]] + then + echo "\e[0;91mFAIL\e[0m: Verilog returned nothing" + #exit 0 + fi + + # Don't care checks + if [[ $veriresult == "x" ]] + then + echo "\e[0;91mFAIL\e[0m: Verilog returned don't cares" + #exit 0 + fi + + # unequal result check + if [[ $cresult -ne $veriresult ]] + then + echo -e "\e[0;91mFAIL\e[0m: Verilog and C output do not match!" + #exit 0 + else + echo -e "\e[0;92mPASS\e[0m" + fi + name=$(echo $benchmark | awk -v FS="/" '{print $NF}') + echo $name","$cycles >> exec.csv +done < benchmark-list-master diff --git a/scripts/script.R b/scripts/script.R new file mode 100644 index 0000000..0be16da --- /dev/null +++ b/scripts/script.R @@ -0,0 +1,29 @@ +library("psych") + +data = read.csv("poly.csv", header=TRUE) +leguptime = (data$legupcycles/data$legupfreqMHz) +veritime = data$vericertcycles/data$vericertfreqMHz +print(lm(veritime ~ leguptime)) +leguputil = data$leguplogicutilisation/427200*100 +veriutil = data$vericertlogicutilisation/427200*100 +print(lm (veriutil ~ leguputil)) +legupct = data$legupcomptime +verict = data$vericertcomptime +print(lm ( verict ~ legupct )) + +cycleslowdown=data$vericertcycles/data$legupcycles + +print("Cycle count slow down") +print(geometric.mean(cycleslowdown)) +print("Wall clock slow down") +print(geometric.mean(veritime/leguptime)) +print("Area overhead") +print(geometric.mean(veriutil/leguputil)) +print("Compilation time speedup") +print(geometric.mean(legupct/verict)) +print("LegUp RAM use") +print(geometric.mean(data$legupregs)) +print("Vericert RAM use") +print(geometric.mean(data$vericertregs)) +print("Area overhead") +print(geometric.mean(data$vericertregs/data$legupregs)) diff --git a/scripts/syn-remote.sh b/scripts/syn-remote.sh new file mode 100755 index 0000000..879db2e --- /dev/null +++ b/scripts/syn-remote.sh @@ -0,0 +1,51 @@ +#! /bin/bash + +#setup +while read benchmark ; +do +echo "Setting up "$benchmark +rm -r $benchmark +mkdir $benchmark +cp $benchmark.v $benchmark/top.v + +done < syn-list + +#synthesis + +count=0 +while read benchmark ; + +do +echo "Synthesising "$benchmark +cd $benchmark +quartus_sh -t ../quartus_synth.tcl & +let "count=count+1" +cd .. + +if [ $count -eq 4 ] +then +echo "I am here" +wait +count=0 +fi + +done < syn-list + +if [ $count -lt 4 ] +then +wait +fi + +#extract +while read benchmark ; do + cd $benchmark + echo $(pwd) + freq=$(grep MHz syn.sta.rpt | tail -2 | head -1 | awk '{print $2}') + lut=$(sed -n -e 8p syn.fit.summary | awk '{print $6}' | sed 's/,//g') + regs=$(sed -n -e 9p syn.fit.summary | awk '{print $4}') + bram=$(sed -n -e 13p syn.fit.summary | awk '{print $5}') + dsp=$(sed -n -e 14p syn.fit.summary | awk '{print $5}') + cd .. + echo $benchmark","$freq","$lut","$regs","$bram","$dsp >> results +done < syn-list + diff --git a/scripts/synthesis-results.scm b/scripts/synthesis-results.scm index 11c66c1..94c169f 100755 --- a/scripts/synthesis-results.scm +++ b/scripts/synthesis-results.scm @@ -91,7 +91,7 @@ (lambda () (list name (xml-matcher (ssax:xml->sxml (current-input-port) '())))))) -(define (to-csv-record b head results) +(define ((to-csv-record b head) results) (let ((res (map (lambda (key) (cadr (assoc key (cadr results)))) head))) (csv:fmt-row (if b res (cons (car results) res))))) @@ -126,8 +126,7 @@ (let ((head (split-at-comma (or (check-opt 'keys) "slice,ramfifo,delay"))) (suppress (split-at-comma (or (check-opt 'suppress) "none"))) (files (get-files-from-op operands))) - (let ((body (map (lambda (f) (to-csv-record (member "name" suppress) head f)) - (convert-files files))) + (let ((body (map (to-csv-record (member "name" suppress) head) (convert-files files))) (header (csv:fmt-row (if (member "name" suppress) head (cons "name" head))))) (with-output (lambda () -- 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 --- .gitignore | 4 ++ benchmarks/polybench-syn-div/exec.csv | 27 ----------- benchmarks/polybench-syn-div/poly.csv | 26 ----------- benchmarks/polybench-syn-div/quartus_synth.tcl | 35 -------------- benchmarks/polybench-syn-div/run-vericert.sh | 41 ----------------- benchmarks/polybench-syn-div/script.R | 29 ------------ benchmarks/polybench-syn-div/setup-syn-vericert.sh | 24 ---------- benchmarks/polybench-syn-div/syn-remote.sh | 51 --------------------- benchmarks/polybench-syn/poly.csv | 26 ----------- benchmarks/polybench-syn/quartus_synth.tcl | 35 -------------- benchmarks/polybench-syn/run-vericert.sh | 46 ------------------- benchmarks/polybench-syn/script.R | 29 ------------ benchmarks/polybench-syn/setup-syn-vericert.sh | 24 ---------- benchmarks/polybench-syn/syn-remote.sh | 51 --------------------- scripts/quartus_synth.tcl | 35 ++++++++++++++ scripts/run-vericert.sh | 46 +++++++++++++++++++ scripts/script.R | 29 ++++++++++++ scripts/syn-remote.sh | 51 +++++++++++++++++++++ scripts/synthesis-results.scm | 5 +- src/hls/main | Bin 0 -> 54136 bytes 20 files changed, 167 insertions(+), 447 deletions(-) delete mode 100644 benchmarks/polybench-syn-div/exec.csv delete mode 100644 benchmarks/polybench-syn-div/poly.csv delete mode 100644 benchmarks/polybench-syn-div/quartus_synth.tcl delete mode 100755 benchmarks/polybench-syn-div/run-vericert.sh delete mode 100644 benchmarks/polybench-syn-div/script.R delete mode 100755 benchmarks/polybench-syn-div/setup-syn-vericert.sh delete mode 100755 benchmarks/polybench-syn-div/syn-remote.sh delete mode 100644 benchmarks/polybench-syn/poly.csv delete mode 100644 benchmarks/polybench-syn/quartus_synth.tcl delete mode 100755 benchmarks/polybench-syn/run-vericert.sh delete mode 100644 benchmarks/polybench-syn/script.R delete mode 100755 benchmarks/polybench-syn/setup-syn-vericert.sh delete mode 100755 benchmarks/polybench-syn/syn-remote.sh create mode 100644 scripts/quartus_synth.tcl create mode 100755 scripts/run-vericert.sh create mode 100644 scripts/script.R create mode 100755 scripts/syn-remote.sh create mode 100755 src/hls/main diff --git a/.gitignore b/.gitignore index 4acfbb0..586f3a7 100644 --- a/.gitignore +++ b/.gitignore @@ -1,3 +1,7 @@ +/benchmarks/**/* +!/benchmarks/**/*.* +!/benchmarks/**/*/ + .*.aux *.a *.cma diff --git a/benchmarks/polybench-syn-div/exec.csv b/benchmarks/polybench-syn-div/exec.csv deleted file mode 100644 index e28109b..0000000 --- a/benchmarks/polybench-syn-div/exec.csv +++ /dev/null @@ -1,27 +0,0 @@ -adi,1422354 -heat-3d,580770 -fdtd-2d,901430 -jacobi-1d,19622 -seidel-2d,664780 -jacobi-2d,344072 -nussinov,954402 -floyd-warshall,5373798 -3mm,536114 -2mm,404478 -doitgen,351988 -bicg,53916 -mvt,70204 -atas,58424 -syrk,271816 -gemver,117394 -symm,240172 -gesummv,37700 -gemm,328104 -trmm,144688 -syr2k,436520 -cholesky,2535686 -trisolv,25192 -lu,2853646 -ludcmp,2601382 -durbin,22974 -covariance,288392 diff --git a/benchmarks/polybench-syn-div/poly.csv b/benchmarks/polybench-syn-div/poly.csv deleted file mode 100644 index fe71983..0000000 --- a/benchmarks/polybench-syn-div/poly.csv +++ /dev/null @@ -1,26 +0,0 @@ -benchmark,legupcycles,legupfreqMHz,leguplogicutilisation,legupregs,leguprams,legupdsps,legupcomptime,vericertcycles,vericertfreqMHz,vericertlogicutilisation,vericertregs,vericertrams,vericertdsps,vericertcomptime -durbin,15106,188.61,2509,4008,0,8,4.77,19852,199.2,3027,6168,0,8,0.077 -lu,482766,244.62,3116,4593,0,10,4.72,2634766,92.97,54806,106037,0,6,0.097 -ludcmp,470843,249.69,3605,5397,0,15,4.87,2401354,83.52,57145,111408,0,10,0.132 -trisolv,35382,213.9,2412,3749,0,3,4.73,33550,112.59,28101,55109,0,2,0.086 -2mm,60088,226.5,1114,1936,0,7,4.80,427098,116.9,32600,63212,0,18,0.136 -3mm,204195,188.96,4210,4801,0,43,4.88,539430,97.13,45073,89719,0,18,0.147 -atas,126288,193.24,3026,3719,0,10,4.80,92000,130.41,28603,56646,0,6,0.069 -bicg,11907,303.4,308,537,0,6,4.78,121134,122.74,30091,58799,0,8,0.097 -mvt,16806,384.47,372,597,0,4,4.79,139194,119.93,30753,60450,0,6,0.130 -doitgen,57199,252.14,909,1402,0,2,5.05,350302,126.18,19898,38461,0,4,0.105 -symm,64903,284.41,2155,3170,0,10,4.63,248930,113.92,27693,54514,0,14,0.115 -syrk,57395,278.01,598,976,0,2,4.73,309018,87.15,76889,57816,0,8,0.094 -syr2k,125705,240.85,3149,3679,0,6,4.85,478040,78.6,120116,82032,0,12,0.116 -trmm,41432,281.61,610,990,0,4,4.71,147528,105.61,64031,40502,0,4,0.089 -gemm,83676,192.68,1029,1544,0,35,4.79,360772,121.61,31853,62126,0,16,0.104 -gemver,28087,303.49,1854,2380,0,8,4.68,175326,107.27,32615,64118,0,14,0.099 -gesummv,6634,310.46,298,504,0,4,4.77,111094,79.97,113876,72908,0,10,0.101 -covariance,109992,245.16,2098,3096,0,5,4.77,297450,110.57,28729,56660,0,4,0.083 -fdtd-2d,214153,262.61,2736,3801,0,2,4.73,831912,108.23,31333,61421,0,6,0.116 -heat-3d,41059,115.54,3132,2910,0,60,4.92,555930,110.42,33915,67273,0,9,0.181 -jacobi-1d,6914,386.25,1355,1885,0,0,4.72,16606,277.93,1636,3305,0,0,0.071 -jacobi-2d,84609,240.79,2347,3185,0,2,4.81,357100,113.53,30393,59782,0,4,0.079 -seidel-2d,345294,232.4,2128,3337,0,2,4.68,875758,127.99,26948,53133,0,2,0.091 -floyd-warshall,1238764,235.52,1869,2367,0,2,4.71,4762766,109.4,59859,118101,0,2,0.094 -nussinov,216467,273.07,1078,1431,0,2,4.79,837958,90.73,60663,119303,0,0,0.080 diff --git a/benchmarks/polybench-syn-div/quartus_synth.tcl b/benchmarks/polybench-syn-div/quartus_synth.tcl deleted file mode 100644 index 6edbf0c..0000000 --- a/benchmarks/polybench-syn-div/quartus_synth.tcl +++ /dev/null @@ -1,35 +0,0 @@ -# PRiME pre-KAPow kernel flow -# Performs pre-KAPow run steps for instrumenting arbitrary Verilog for power monitoring -# James Davis, 2015 - -load_package flow - -project_new -overwrite syn -set_global_assignment -name FAMILY "Arria 10" -set_global_assignment -name DEVICE 10AX115H4F34E3LG -set_global_assignment -name SYSTEMVERILOG_FILE top.v -set_global_assignment -name TOP_LEVEL_ENTITY main -#set_global_assignment -name SDC_FILE syn.sdc -#set_global_assignment -name auto_resource_sharing on -#set_global_assignment -name enable_state_machine_inference on -#set_global_assignment -name optimization_technique area -#set_global_assignment -name synthesis_effort fast -#set_global_assignment -name AUTO_RAM_RECOGNITION on -#set_global_assignment -name remove_duplicate_registers on -#set_instance_assignment -name RAMSTYLE_ATTRIBUTE LOGIC -to ram - -execute_module -tool map - -execute_module -tool fit - -execute_module -tool sta - -#execute_module -tool eda -args "--simulation --tool=vcs" - -# set_global_assignment -name POWER_OUTPUT_SAF_NAME ${kernel}.asf -# set_global_assignment -name POWER_DEFAULT_INPUT_IO_TOGGLE_RATE "12.5 %" -# set_global_assignment -name POWER_REPORT_SIGNAL_ACTIVITY ON -# set_global_assignment -name POWER_REPORT_POWER_DISSIPATION ON -# execute_module -tool pow - -project_close diff --git a/benchmarks/polybench-syn-div/run-vericert.sh b/benchmarks/polybench-syn-div/run-vericert.sh deleted file mode 100755 index 6cf4cd9..0000000 --- a/benchmarks/polybench-syn-div/run-vericert.sh +++ /dev/null @@ -1,41 +0,0 @@ -#!/usr/bin/env bash - -rm exec.csv - -top=$(pwd) - #set up -while read benchmark ; do - echo "Running "$benchmark - ./$benchmark.gcc > $benchmark.clog - cresult=$(cat $benchmark.clog | cut -d' ' -f2) - echo "C output: "$cresult - ./$benchmark.iver > $benchmark.tmp - veriresult=$(tail -1 $benchmark.tmp | cut -d' ' -f2) - cycles=$(tail -2 $benchmark.tmp | head -1 | tr -s ' ' | cut -d' ' -f2) - echo "Verilog output: "$veriresult - - #Undefined checks - if test -z $veriresult - then - echo "FAIL: Verilog returned nothing" - #exit 0 - fi - - # Don't care checks - if [ $veriresult == "x" ] - then - echo "FAIL: Verilog returned don't cares" - #exit 0 - fi - - # unequal result check - if [ $cresult -ne $veriresult ] - then - echo "FAIL: Verilog and C output do not match!" - #exit 0 - else - echo "PASS" - fi - name=$(echo $benchmark | awk -v FS="/" '{print $NF}') - echo $name","$cycles >> exec.csv -done < benchmark-list-master diff --git a/benchmarks/polybench-syn-div/script.R b/benchmarks/polybench-syn-div/script.R deleted file mode 100644 index 0be16da..0000000 --- a/benchmarks/polybench-syn-div/script.R +++ /dev/null @@ -1,29 +0,0 @@ -library("psych") - -data = read.csv("poly.csv", header=TRUE) -leguptime = (data$legupcycles/data$legupfreqMHz) -veritime = data$vericertcycles/data$vericertfreqMHz -print(lm(veritime ~ leguptime)) -leguputil = data$leguplogicutilisation/427200*100 -veriutil = data$vericertlogicutilisation/427200*100 -print(lm (veriutil ~ leguputil)) -legupct = data$legupcomptime -verict = data$vericertcomptime -print(lm ( verict ~ legupct )) - -cycleslowdown=data$vericertcycles/data$legupcycles - -print("Cycle count slow down") -print(geometric.mean(cycleslowdown)) -print("Wall clock slow down") -print(geometric.mean(veritime/leguptime)) -print("Area overhead") -print(geometric.mean(veriutil/leguputil)) -print("Compilation time speedup") -print(geometric.mean(legupct/verict)) -print("LegUp RAM use") -print(geometric.mean(data$legupregs)) -print("Vericert RAM use") -print(geometric.mean(data$vericertregs)) -print("Area overhead") -print(geometric.mean(data$vericertregs/data$legupregs)) diff --git a/benchmarks/polybench-syn-div/setup-syn-vericert.sh b/benchmarks/polybench-syn-div/setup-syn-vericert.sh deleted file mode 100755 index 22356f7..0000000 --- a/benchmarks/polybench-syn-div/setup-syn-vericert.sh +++ /dev/null @@ -1,24 +0,0 @@ -#! /bin/bash - -top=$(pwd) - #set up - basedir=poly-syn - sshhost=$1 - ssh $sshhost "cd ~; rm -r $basedir" - ssh $sshhost "cd ~; mkdir $basedir" - scp quartus_synth.tcl $sshhost:$basedir - scp syn-remote.sh $sshhost:$basedir - rm syn-list - - while read benchmark ; - do - echo "Copying "$benchmark" over" - name=$(echo $benchmark | awk -v FS="/" '{print $NF}') - echo "Name: "$name - benchdir="~/$basedir/$name" - scp $benchmark.v $sshhost:~/$basedir - echo $name >> syn-list - done < benchmark-list-master - - # copy list over - scp syn-list $sshhost:$basedir diff --git a/benchmarks/polybench-syn-div/syn-remote.sh b/benchmarks/polybench-syn-div/syn-remote.sh deleted file mode 100755 index 879db2e..0000000 --- a/benchmarks/polybench-syn-div/syn-remote.sh +++ /dev/null @@ -1,51 +0,0 @@ -#! /bin/bash - -#setup -while read benchmark ; -do -echo "Setting up "$benchmark -rm -r $benchmark -mkdir $benchmark -cp $benchmark.v $benchmark/top.v - -done < syn-list - -#synthesis - -count=0 -while read benchmark ; - -do -echo "Synthesising "$benchmark -cd $benchmark -quartus_sh -t ../quartus_synth.tcl & -let "count=count+1" -cd .. - -if [ $count -eq 4 ] -then -echo "I am here" -wait -count=0 -fi - -done < syn-list - -if [ $count -lt 4 ] -then -wait -fi - -#extract -while read benchmark ; do - cd $benchmark - echo $(pwd) - freq=$(grep MHz syn.sta.rpt | tail -2 | head -1 | awk '{print $2}') - lut=$(sed -n -e 8p syn.fit.summary | awk '{print $6}' | sed 's/,//g') - regs=$(sed -n -e 9p syn.fit.summary | awk '{print $4}') - bram=$(sed -n -e 13p syn.fit.summary | awk '{print $5}') - dsp=$(sed -n -e 14p syn.fit.summary | awk '{print $5}') - cd .. - echo $benchmark","$freq","$lut","$regs","$bram","$dsp >> results -done < syn-list - diff --git a/benchmarks/polybench-syn/poly.csv b/benchmarks/polybench-syn/poly.csv deleted file mode 100644 index fe71983..0000000 --- a/benchmarks/polybench-syn/poly.csv +++ /dev/null @@ -1,26 +0,0 @@ -benchmark,legupcycles,legupfreqMHz,leguplogicutilisation,legupregs,leguprams,legupdsps,legupcomptime,vericertcycles,vericertfreqMHz,vericertlogicutilisation,vericertregs,vericertrams,vericertdsps,vericertcomptime -durbin,15106,188.61,2509,4008,0,8,4.77,19852,199.2,3027,6168,0,8,0.077 -lu,482766,244.62,3116,4593,0,10,4.72,2634766,92.97,54806,106037,0,6,0.097 -ludcmp,470843,249.69,3605,5397,0,15,4.87,2401354,83.52,57145,111408,0,10,0.132 -trisolv,35382,213.9,2412,3749,0,3,4.73,33550,112.59,28101,55109,0,2,0.086 -2mm,60088,226.5,1114,1936,0,7,4.80,427098,116.9,32600,63212,0,18,0.136 -3mm,204195,188.96,4210,4801,0,43,4.88,539430,97.13,45073,89719,0,18,0.147 -atas,126288,193.24,3026,3719,0,10,4.80,92000,130.41,28603,56646,0,6,0.069 -bicg,11907,303.4,308,537,0,6,4.78,121134,122.74,30091,58799,0,8,0.097 -mvt,16806,384.47,372,597,0,4,4.79,139194,119.93,30753,60450,0,6,0.130 -doitgen,57199,252.14,909,1402,0,2,5.05,350302,126.18,19898,38461,0,4,0.105 -symm,64903,284.41,2155,3170,0,10,4.63,248930,113.92,27693,54514,0,14,0.115 -syrk,57395,278.01,598,976,0,2,4.73,309018,87.15,76889,57816,0,8,0.094 -syr2k,125705,240.85,3149,3679,0,6,4.85,478040,78.6,120116,82032,0,12,0.116 -trmm,41432,281.61,610,990,0,4,4.71,147528,105.61,64031,40502,0,4,0.089 -gemm,83676,192.68,1029,1544,0,35,4.79,360772,121.61,31853,62126,0,16,0.104 -gemver,28087,303.49,1854,2380,0,8,4.68,175326,107.27,32615,64118,0,14,0.099 -gesummv,6634,310.46,298,504,0,4,4.77,111094,79.97,113876,72908,0,10,0.101 -covariance,109992,245.16,2098,3096,0,5,4.77,297450,110.57,28729,56660,0,4,0.083 -fdtd-2d,214153,262.61,2736,3801,0,2,4.73,831912,108.23,31333,61421,0,6,0.116 -heat-3d,41059,115.54,3132,2910,0,60,4.92,555930,110.42,33915,67273,0,9,0.181 -jacobi-1d,6914,386.25,1355,1885,0,0,4.72,16606,277.93,1636,3305,0,0,0.071 -jacobi-2d,84609,240.79,2347,3185,0,2,4.81,357100,113.53,30393,59782,0,4,0.079 -seidel-2d,345294,232.4,2128,3337,0,2,4.68,875758,127.99,26948,53133,0,2,0.091 -floyd-warshall,1238764,235.52,1869,2367,0,2,4.71,4762766,109.4,59859,118101,0,2,0.094 -nussinov,216467,273.07,1078,1431,0,2,4.79,837958,90.73,60663,119303,0,0,0.080 diff --git a/benchmarks/polybench-syn/quartus_synth.tcl b/benchmarks/polybench-syn/quartus_synth.tcl deleted file mode 100644 index 6edbf0c..0000000 --- a/benchmarks/polybench-syn/quartus_synth.tcl +++ /dev/null @@ -1,35 +0,0 @@ -# PRiME pre-KAPow kernel flow -# Performs pre-KAPow run steps for instrumenting arbitrary Verilog for power monitoring -# James Davis, 2015 - -load_package flow - -project_new -overwrite syn -set_global_assignment -name FAMILY "Arria 10" -set_global_assignment -name DEVICE 10AX115H4F34E3LG -set_global_assignment -name SYSTEMVERILOG_FILE top.v -set_global_assignment -name TOP_LEVEL_ENTITY main -#set_global_assignment -name SDC_FILE syn.sdc -#set_global_assignment -name auto_resource_sharing on -#set_global_assignment -name enable_state_machine_inference on -#set_global_assignment -name optimization_technique area -#set_global_assignment -name synthesis_effort fast -#set_global_assignment -name AUTO_RAM_RECOGNITION on -#set_global_assignment -name remove_duplicate_registers on -#set_instance_assignment -name RAMSTYLE_ATTRIBUTE LOGIC -to ram - -execute_module -tool map - -execute_module -tool fit - -execute_module -tool sta - -#execute_module -tool eda -args "--simulation --tool=vcs" - -# set_global_assignment -name POWER_OUTPUT_SAF_NAME ${kernel}.asf -# set_global_assignment -name POWER_DEFAULT_INPUT_IO_TOGGLE_RATE "12.5 %" -# set_global_assignment -name POWER_REPORT_SIGNAL_ACTIVITY ON -# set_global_assignment -name POWER_REPORT_POWER_DISSIPATION ON -# execute_module -tool pow - -project_close diff --git a/benchmarks/polybench-syn/run-vericert.sh b/benchmarks/polybench-syn/run-vericert.sh deleted file mode 100755 index 9deaa10..0000000 --- a/benchmarks/polybench-syn/run-vericert.sh +++ /dev/null @@ -1,46 +0,0 @@ -#!/usr/bin/env bash - -rm exec.csv - -top=$(pwd) -#set up -while read benchmark ; do - printf "%10s\t" $(echo "$benchmark" | sed -e 's|/| |g') - ./$benchmark.gcc > $benchmark.clog - cresult=$(cat $benchmark.clog | cut -d' ' -f2) - #echo "C output: "$cresult - #./$benchmark.iver > $benchmark.tmp - if [[ ! -f ./$benchmark.verilator/Vmain ]]; then - echo -e "\e[0;91mFAIL\e[0m: Verilog failed compilation" - continue - fi - ./$benchmark.verilator/Vmain > $benchmark.tmp - veriresult=$(tail -1 $benchmark.tmp | cut -d' ' -f2) - cycles=$(tail -2 $benchmark.tmp | head -1 | tr -s ' ' | cut -d' ' -f2) - #echo "Verilog output: "$veriresult - - #Undefined checks - if [[ -z "$veriresult" ]] - then - echo "\e[0;91mFAIL\e[0m: Verilog returned nothing" - #exit 0 - fi - - # Don't care checks - if [[ $veriresult == "x" ]] - then - echo "\e[0;91mFAIL\e[0m: Verilog returned don't cares" - #exit 0 - fi - - # unequal result check - if [[ $cresult -ne $veriresult ]] - then - echo -e "\e[0;91mFAIL\e[0m: Verilog and C output do not match!" - #exit 0 - else - echo -e "\e[0;92mPASS\e[0m" - fi - name=$(echo $benchmark | awk -v FS="/" '{print $NF}') - echo $name","$cycles >> exec.csv -done < benchmark-list-master diff --git a/benchmarks/polybench-syn/script.R b/benchmarks/polybench-syn/script.R deleted file mode 100644 index 0be16da..0000000 --- a/benchmarks/polybench-syn/script.R +++ /dev/null @@ -1,29 +0,0 @@ -library("psych") - -data = read.csv("poly.csv", header=TRUE) -leguptime = (data$legupcycles/data$legupfreqMHz) -veritime = data$vericertcycles/data$vericertfreqMHz -print(lm(veritime ~ leguptime)) -leguputil = data$leguplogicutilisation/427200*100 -veriutil = data$vericertlogicutilisation/427200*100 -print(lm (veriutil ~ leguputil)) -legupct = data$legupcomptime -verict = data$vericertcomptime -print(lm ( verict ~ legupct )) - -cycleslowdown=data$vericertcycles/data$legupcycles - -print("Cycle count slow down") -print(geometric.mean(cycleslowdown)) -print("Wall clock slow down") -print(geometric.mean(veritime/leguptime)) -print("Area overhead") -print(geometric.mean(veriutil/leguputil)) -print("Compilation time speedup") -print(geometric.mean(legupct/verict)) -print("LegUp RAM use") -print(geometric.mean(data$legupregs)) -print("Vericert RAM use") -print(geometric.mean(data$vericertregs)) -print("Area overhead") -print(geometric.mean(data$vericertregs/data$legupregs)) diff --git a/benchmarks/polybench-syn/setup-syn-vericert.sh b/benchmarks/polybench-syn/setup-syn-vericert.sh deleted file mode 100755 index 22356f7..0000000 --- a/benchmarks/polybench-syn/setup-syn-vericert.sh +++ /dev/null @@ -1,24 +0,0 @@ -#! /bin/bash - -top=$(pwd) - #set up - basedir=poly-syn - sshhost=$1 - ssh $sshhost "cd ~; rm -r $basedir" - ssh $sshhost "cd ~; mkdir $basedir" - scp quartus_synth.tcl $sshhost:$basedir - scp syn-remote.sh $sshhost:$basedir - rm syn-list - - while read benchmark ; - do - echo "Copying "$benchmark" over" - name=$(echo $benchmark | awk -v FS="/" '{print $NF}') - echo "Name: "$name - benchdir="~/$basedir/$name" - scp $benchmark.v $sshhost:~/$basedir - echo $name >> syn-list - done < benchmark-list-master - - # copy list over - scp syn-list $sshhost:$basedir diff --git a/benchmarks/polybench-syn/syn-remote.sh b/benchmarks/polybench-syn/syn-remote.sh deleted file mode 100755 index 879db2e..0000000 --- a/benchmarks/polybench-syn/syn-remote.sh +++ /dev/null @@ -1,51 +0,0 @@ -#! /bin/bash - -#setup -while read benchmark ; -do -echo "Setting up "$benchmark -rm -r $benchmark -mkdir $benchmark -cp $benchmark.v $benchmark/top.v - -done < syn-list - -#synthesis - -count=0 -while read benchmark ; - -do -echo "Synthesising "$benchmark -cd $benchmark -quartus_sh -t ../quartus_synth.tcl & -let "count=count+1" -cd .. - -if [ $count -eq 4 ] -then -echo "I am here" -wait -count=0 -fi - -done < syn-list - -if [ $count -lt 4 ] -then -wait -fi - -#extract -while read benchmark ; do - cd $benchmark - echo $(pwd) - freq=$(grep MHz syn.sta.rpt | tail -2 | head -1 | awk '{print $2}') - lut=$(sed -n -e 8p syn.fit.summary | awk '{print $6}' | sed 's/,//g') - regs=$(sed -n -e 9p syn.fit.summary | awk '{print $4}') - bram=$(sed -n -e 13p syn.fit.summary | awk '{print $5}') - dsp=$(sed -n -e 14p syn.fit.summary | awk '{print $5}') - cd .. - echo $benchmark","$freq","$lut","$regs","$bram","$dsp >> results -done < syn-list - diff --git a/scripts/quartus_synth.tcl b/scripts/quartus_synth.tcl new file mode 100644 index 0000000..6edbf0c --- /dev/null +++ b/scripts/quartus_synth.tcl @@ -0,0 +1,35 @@ +# PRiME pre-KAPow kernel flow +# Performs pre-KAPow run steps for instrumenting arbitrary Verilog for power monitoring +# James Davis, 2015 + +load_package flow + +project_new -overwrite syn +set_global_assignment -name FAMILY "Arria 10" +set_global_assignment -name DEVICE 10AX115H4F34E3LG +set_global_assignment -name SYSTEMVERILOG_FILE top.v +set_global_assignment -name TOP_LEVEL_ENTITY main +#set_global_assignment -name SDC_FILE syn.sdc +#set_global_assignment -name auto_resource_sharing on +#set_global_assignment -name enable_state_machine_inference on +#set_global_assignment -name optimization_technique area +#set_global_assignment -name synthesis_effort fast +#set_global_assignment -name AUTO_RAM_RECOGNITION on +#set_global_assignment -name remove_duplicate_registers on +#set_instance_assignment -name RAMSTYLE_ATTRIBUTE LOGIC -to ram + +execute_module -tool map + +execute_module -tool fit + +execute_module -tool sta + +#execute_module -tool eda -args "--simulation --tool=vcs" + +# set_global_assignment -name POWER_OUTPUT_SAF_NAME ${kernel}.asf +# set_global_assignment -name POWER_DEFAULT_INPUT_IO_TOGGLE_RATE "12.5 %" +# set_global_assignment -name POWER_REPORT_SIGNAL_ACTIVITY ON +# set_global_assignment -name POWER_REPORT_POWER_DISSIPATION ON +# execute_module -tool pow + +project_close diff --git a/scripts/run-vericert.sh b/scripts/run-vericert.sh new file mode 100755 index 0000000..9deaa10 --- /dev/null +++ b/scripts/run-vericert.sh @@ -0,0 +1,46 @@ +#!/usr/bin/env bash + +rm exec.csv + +top=$(pwd) +#set up +while read benchmark ; do + printf "%10s\t" $(echo "$benchmark" | sed -e 's|/| |g') + ./$benchmark.gcc > $benchmark.clog + cresult=$(cat $benchmark.clog | cut -d' ' -f2) + #echo "C output: "$cresult + #./$benchmark.iver > $benchmark.tmp + if [[ ! -f ./$benchmark.verilator/Vmain ]]; then + echo -e "\e[0;91mFAIL\e[0m: Verilog failed compilation" + continue + fi + ./$benchmark.verilator/Vmain > $benchmark.tmp + veriresult=$(tail -1 $benchmark.tmp | cut -d' ' -f2) + cycles=$(tail -2 $benchmark.tmp | head -1 | tr -s ' ' | cut -d' ' -f2) + #echo "Verilog output: "$veriresult + + #Undefined checks + if [[ -z "$veriresult" ]] + then + echo "\e[0;91mFAIL\e[0m: Verilog returned nothing" + #exit 0 + fi + + # Don't care checks + if [[ $veriresult == "x" ]] + then + echo "\e[0;91mFAIL\e[0m: Verilog returned don't cares" + #exit 0 + fi + + # unequal result check + if [[ $cresult -ne $veriresult ]] + then + echo -e "\e[0;91mFAIL\e[0m: Verilog and C output do not match!" + #exit 0 + else + echo -e "\e[0;92mPASS\e[0m" + fi + name=$(echo $benchmark | awk -v FS="/" '{print $NF}') + echo $name","$cycles >> exec.csv +done < benchmark-list-master diff --git a/scripts/script.R b/scripts/script.R new file mode 100644 index 0000000..0be16da --- /dev/null +++ b/scripts/script.R @@ -0,0 +1,29 @@ +library("psych") + +data = read.csv("poly.csv", header=TRUE) +leguptime = (data$legupcycles/data$legupfreqMHz) +veritime = data$vericertcycles/data$vericertfreqMHz +print(lm(veritime ~ leguptime)) +leguputil = data$leguplogicutilisation/427200*100 +veriutil = data$vericertlogicutilisation/427200*100 +print(lm (veriutil ~ leguputil)) +legupct = data$legupcomptime +verict = data$vericertcomptime +print(lm ( verict ~ legupct )) + +cycleslowdown=data$vericertcycles/data$legupcycles + +print("Cycle count slow down") +print(geometric.mean(cycleslowdown)) +print("Wall clock slow down") +print(geometric.mean(veritime/leguptime)) +print("Area overhead") +print(geometric.mean(veriutil/leguputil)) +print("Compilation time speedup") +print(geometric.mean(legupct/verict)) +print("LegUp RAM use") +print(geometric.mean(data$legupregs)) +print("Vericert RAM use") +print(geometric.mean(data$vericertregs)) +print("Area overhead") +print(geometric.mean(data$vericertregs/data$legupregs)) diff --git a/scripts/syn-remote.sh b/scripts/syn-remote.sh new file mode 100755 index 0000000..879db2e --- /dev/null +++ b/scripts/syn-remote.sh @@ -0,0 +1,51 @@ +#! /bin/bash + +#setup +while read benchmark ; +do +echo "Setting up "$benchmark +rm -r $benchmark +mkdir $benchmark +cp $benchmark.v $benchmark/top.v + +done < syn-list + +#synthesis + +count=0 +while read benchmark ; + +do +echo "Synthesising "$benchmark +cd $benchmark +quartus_sh -t ../quartus_synth.tcl & +let "count=count+1" +cd .. + +if [ $count -eq 4 ] +then +echo "I am here" +wait +count=0 +fi + +done < syn-list + +if [ $count -lt 4 ] +then +wait +fi + +#extract +while read benchmark ; do + cd $benchmark + echo $(pwd) + freq=$(grep MHz syn.sta.rpt | tail -2 | head -1 | awk '{print $2}') + lut=$(sed -n -e 8p syn.fit.summary | awk '{print $6}' | sed 's/,//g') + regs=$(sed -n -e 9p syn.fit.summary | awk '{print $4}') + bram=$(sed -n -e 13p syn.fit.summary | awk '{print $5}') + dsp=$(sed -n -e 14p syn.fit.summary | awk '{print $5}') + cd .. + echo $benchmark","$freq","$lut","$regs","$bram","$dsp >> results +done < syn-list + diff --git a/scripts/synthesis-results.scm b/scripts/synthesis-results.scm index 11c66c1..94c169f 100755 --- a/scripts/synthesis-results.scm +++ b/scripts/synthesis-results.scm @@ -91,7 +91,7 @@ (lambda () (list name (xml-matcher (ssax:xml->sxml (current-input-port) '())))))) -(define (to-csv-record b head results) +(define ((to-csv-record b head) results) (let ((res (map (lambda (key) (cadr (assoc key (cadr results)))) head))) (csv:fmt-row (if b res (cons (car results) res))))) @@ -126,8 +126,7 @@ (let ((head (split-at-comma (or (check-opt 'keys) "slice,ramfifo,delay"))) (suppress (split-at-comma (or (check-opt 'suppress) "none"))) (files (get-files-from-op operands))) - (let ((body (map (lambda (f) (to-csv-record (member "name" suppress) head f)) - (convert-files files))) + (let ((body (map (to-csv-record (member "name" suppress) head) (convert-files files))) (header (csv:fmt-row (if (member "name" suppress) head (cons "name" head))))) (with-output (lambda () 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 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 0580932cb0da7fac6b8aa5f5b98beb824a3fff50 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 16:42:04 +0000 Subject: Fix Makefiles in build and for benchmarks --- Makefile | 4 ++++ benchmarks/polybench-syn/common.mk | 8 +++++--- benchmarks/polybench-syn/stencils/Makefile | 1 + dune | 2 +- scripts/Makefile | 7 +++++-- test/Makefile | 2 +- 6 files changed, 17 insertions(+), 7 deletions(-) diff --git a/Makefile b/Makefile index e9df0dd..5e1ae44 100644 --- a/Makefile +++ b/Makefile @@ -22,6 +22,7 @@ COQMAKE := $(COQBIN)coq_makefile COQDOCFLAGS := --no-lib-name -l VS := src/Compiler.v src/Simulator.v src/HLSOpts.v $(foreach d, common hls bourdoncle, src/$(d)/*.v) +LIT := $(wildcard lit/*.org) PREFIX ?= . @@ -85,6 +86,9 @@ clean:: Makefile.coq $(MAKE) -C test clean rm -f Makefile.coq +detangle-all: + $(foreach vs,$(VS),emacs --batch --find-file $(vs) --eval "(progn(require 'org)(require 'ob-tangle)(org-babel-detangle))") + clean:: rm -f */*.v.d */*.glob */*.vo */*~ *~ rm -f src/extraction/*.ml src/extraction/*.mli diff --git a/benchmarks/polybench-syn/common.mk b/benchmarks/polybench-syn/common.mk index 8e67294..0ebc527 100644 --- a/benchmarks/polybench-syn/common.mk +++ b/benchmarks/polybench-syn/common.mk @@ -1,15 +1,16 @@ VERICERT ?= vericert -VERICERT_OPTS ?= -DSYNTHESIS -fschedule -fif-conv +VERICERT_OPTS ?= -DSYNTHESIS -fschedule IVERILOG ?= iverilog IVERILOG_OPTS ?= VERILATOR ?= verilator -VERILATOR_OPTS ?= -Wno-fatal --top main --exe /home/ymherklotz/projects/vericert/scripts/verilator_main.cpp +VERILATOR_OPTS ?= -Wno-fatal -Wno-lint -Wno-style -Wno-WIDTH --top main --exe /home/ymherklotz/projects/vericert/scripts/verilator_main.cpp TARGETS ?= %.v: %.c + @echo -e "\033[0;35mMAKE\033[0m" $< $(VERICERT) $(VERICERT_OPTS) $< -o $@ %.iver: %.v @@ -20,7 +21,8 @@ TARGETS ?= %.verilator: %.v $(VERILATOR) $(VERILATOR_OPTS) --Mdir $@ --cc $< - $(MAKE) -C $@ -f Vmain.mk + @echo -e $(MAKE) -C $@ -f Vmain.mk + @$(MAKE) -C $@ -f Vmain.mk &>/dev/null %: %.iver %.gcc %.verilator cp $< $@ diff --git a/benchmarks/polybench-syn/stencils/Makefile b/benchmarks/polybench-syn/stencils/Makefile index d2e1c9b..9b89b0c 100644 --- a/benchmarks/polybench-syn/stencils/Makefile +++ b/benchmarks/polybench-syn/stencils/Makefile @@ -3,4 +3,5 @@ TARGETS := adi fdtd-2d heat-3d jacobi-1d jacobi-2d seidel-2d include ../common.mk adi.v: adi.c + @echo -e "\033[0;35mMAKE\033[0m" $< $(VERICERT) $(VERICERT_OPTS) -O0 -finline $< -o $@ diff --git a/dune b/dune index f58c64e..9bd70aa 100644 --- a/dune +++ b/dune @@ -8,4 +8,4 @@ (public_name vericert) (modules_without_implementation c debugTypes dwarfTypes) (libraries menhirLib str unix ocamlgraph) - (flags (:standard -warn-error -A))) + (flags (:standard -warn-error -A -w -8-9-16-20-26-27-32..36-39-41-44..45-50-60-67))) diff --git a/scripts/Makefile b/scripts/Makefile index 0b5ff33..3892723 100644 --- a/scripts/Makefile +++ b/scripts/Makefile @@ -1,9 +1,12 @@ PREFIX ?= .. -all: synthesis +all: synthesis-results %: %.scm - chicken-csc -static -o $@ $< + $(eval TMP := $(shell mktemp)) + echo "(main (command-line-arguments))" >$(TMP) + chicken-csc -static -epilogue $(TMP) -output-file $@ $< + rm $(TMP) %.1: %.org emacs --batch --file $< --load ../docs/res/publish.el --funcall org-man-export-to-man diff --git a/test/Makefile b/test/Makefile index fa482c6..aa76e70 100644 --- a/test/Makefile +++ b/test/Makefile @@ -18,7 +18,7 @@ all: $(TESTS) @$(CC) $(CFLAGS) -o $@ $< %.v: %.c - @$(VERICERT) $(VERICERT_OPTS) -o $@ $< + $(VERICERT) $(VERICERT_OPTS) -o $@ $< %.iver: %.v @$(IVERILOG) $(IVERILOG_OPTS) -o $@ -- $< -- 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 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(-) 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/common/Vericertlib.v | 43 +++++++++++++++++++++++++++++++++++++++++++ src/hls/Abstr.v | 42 ------------------------------------------ 2 files changed, 43 insertions(+), 42 deletions(-) diff --git a/src/common/Vericertlib.v b/src/common/Vericertlib.v index 389a74f..5540f34 100644 --- a/src/common/Vericertlib.v +++ b/src/common/Vericertlib.v @@ -28,6 +28,7 @@ Require Import Coq.micromega.Lia. Require Export compcert.lib.Coqlib. Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. Require Import vericert.common.Show. @@ -243,3 +244,45 @@ Definition debug_show {A B : Type} `{Show A} (a : A) (b : B) : B := Definition debug_show_msg {A B : Type} `{Show A} (s : string) (a : A) (b : B) : B := let unused := debug_print (s ++ show a) in b. + +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. 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 533268faef6b566aba693234a9005172ffa8f494 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 16:43:06 +0000 Subject: Add to benchmark script --- scripts/synthesis-results.scm | 106 ++++++++++++++++++++++++++++-------------- 1 file changed, 70 insertions(+), 36 deletions(-) diff --git a/scripts/synthesis-results.scm b/scripts/synthesis-results.scm index 94c169f..b1a7349 100755 --- a/scripts/synthesis-results.scm +++ b/scripts/synthesis-results.scm @@ -1,4 +1,4 @@ -#! /usr/bin/chicken-csi -s +#! /usr/bin/chicken-csi -ss ;; -*- mode: scheme -*- ;; ;; Copyright (C) 2022 Yann Herklotz @@ -16,11 +16,13 @@ ;; You should have received a copy of the GNU General Public License ;; along with this program. If not, see . -(import (chicken port) +(import (chicken file) + (chicken io) + (chicken irregex) + (chicken port) (chicken process-context) + (chicken sort) (chicken string) - (chicken irregex) - (chicken file) args matchable srfi-193 @@ -32,32 +34,34 @@ (define (check-opt b) (cdr (or (assoc b options) `(,b . #f)))) (define opts - (list (args:make-option (v verbose) (optional: "LEVEL") "Debug level [default: 0]" - (set! arg (or arg "0"))) + (list (args:make-option (v verbose) (optional: "LEVEL") "Debug level [default: 0]" + (set! arg (or arg "0"))) (args:make-option (k keys) (required: "KEY,KEY,...") "Keys to display [default: slice,ramfifo,delay]") (args:make-option (o output) (required: "FILE") "Output file") - (args:make-option (s suppress) (required: "TYPE,TYPE,...") "Values to suppress from output [default: none]") - (args:make-option (x xml) #:none "Output raw XML from the synthesis tool") + (args:make-option (default-sort) #:none "Don't the names of the benchmarks") + (args:make-option (suppress) (required: "TYPE,TYPE,...") "Values to suppress from output [default: none]") (args:make-option (c csv) #:none "Output processed CSV") - (args:make-option (V version) #:none "Display version" - (print "synthesise v0.1.0") - (exit)) - (args:make-option (h help) #:none "Display this text" - (usage)))) + (args:make-option (cycle-file) (required: "FILE") "File which contains the cycle counts for the benchmark") + (args:make-option (org) #:none "Output processed Org") + (args:make-option (V version) #:none "Display version" + (print "synthesis-results v0.2.0") + (exit)) + (args:make-option (h help) #:none "Display this text" + (usage)))) (: description string) (define description "synthesis-results: sends a verilog file to be synthesised and returns results as a CSV file.") (define (usage) - (with-output-to-port (current-error-port) - (lambda () + (with-output-to-port (current-error-port) + (lambda () (print description) (newline) - (print "Usage: " (program-name) " [options...] [files...]") - (newline) - (print (args:usage opts)) - (print "Report bugs to git at yannherklotz dot com."))) + (print "Usage: " (program-name) " [options...] [files...]") + (newline) + (print (args:usage opts)) + (print "Report bugs to git at yannherklotz dot com."))) (exit 1)) (define (map-names n) @@ -75,6 +79,8 @@ (define (csv:fmt-row l) (string-intersperse (map ->string l) ",")) +(define (org:fmt-row l) (string-append "| " (string-intersperse (map ->string l) " | ") " |")) + (define (csv:fmt-table-string l) (apply string-append (map (lambda (s) (string-append s "\n")) l))) (define (csv:fmt-table l) (apply string-append (map (lambda (s) (string-append s "\n")) @@ -84,28 +90,40 @@ (match xml [('*TOP* _ ('document ('application ('section _ . r)))) (map (match-lambda - [('item ('@ ('value v) ('stringID s))) (list (map-names s) (string->number v))]) r)])) + [('item ('@ ('value v) ('stringID s))) (cons (map-names s) (string->number v))]) r)])) (define (parse-xml name file) (with-input-from-file file (lambda () - (list name (xml-matcher (ssax:xml->sxml (current-input-port) '())))))) + (cons name (xml-matcher (ssax:xml->sxml (current-input-port) '())))))) + +(define (ifn-cons b c l) + (if b l (cons c l))) -(define ((to-csv-record b head) results) +(define ((to-csv-record fmt-row b head) results) (let ((res (map (lambda (key) - (cadr (assoc key (cadr results)))) head))) - (csv:fmt-row (if b res (cons (car results) res))))) + (cdr (assoc key (cadr results)))) head))) + (fmt-row (ifn-cons b (car results) res)))) (: path-to-name (string -> string)) (define (path-to-name path) (irregex-replace "^.*?([^/]+)_report\\.xml$" path 1)) -(define (convert-files files) - (map (lambda (f) (parse-xml (path-to-name f) f)) files)) +(define (order-data d1 d2) + (match (list d1 d2) + [((n1 . _) (n2 . _)) (string (list-of string))) (define (split-at-comma s) (string-split s ",")) +(define (parse-cycles f) + (if f (with-input-from-file f + (lambda () (map (match-lambda [(a b) (cons a b)]) (map split-at-comma (read-lines))))) #f)) + (: find-all-xml (string -> (list-of string))) (define (find-all-xml dir) (find-files dir #:test ".*\\.xml$")) @@ -120,17 +138,33 @@ (with-output-to-file (check-opt 'output) thk) (thk))) +(define (row-fmt) + (if (check-opt 'org) org:fmt-row csv:fmt-row)) + +(define (cycle-file operands) + (let ((dir (match operands + [(d) (when (directory-exists? d) d)] + [_ #f]))) + (or (check-opt 'cycle-file) (if dir (string-append dir "/exec.csv") #f)))) + +(define (add-cycles data cycles) + (if cycles + (map (lambda (f) (list (car f) (cons (cons "cycles" (cdr (assoc (car f) cycles))) (cdr f)))) data) + data)) + (define (main args) (set!-values (options operands) - (args:parse args opts)) - (let ((head (split-at-comma (or (check-opt 'keys) "slice,ramfifo,delay"))) + (args:parse args opts)) + (let ((head (split-at-comma (or (check-opt 'keys) "slice,ramfifo,delay,cycles"))) (suppress (split-at-comma (or (check-opt 'suppress) "none"))) (files (get-files-from-op operands))) - (let ((body (map (to-csv-record (member "name" suppress) head) (convert-files files))) - (header (csv:fmt-row (if (member "name" suppress) head (cons "name" head))))) - (with-output - (lambda () - (display (csv:fmt-table-string - (if (member "header" suppress) body (cons header body))))))))) - -(main (command-line-arguments)) + (let ((data (convert-files (not (check-opt 'default-sort)) files)) + (header ((row-fmt) (ifn-cons (member "name" suppress) "name" head))) + (cycles (parse-cycles (cycle-file operands)))) + (with-output + (lambda () + (display (csv:fmt-table-string + (ifn-cons (member "header" suppress) header + (map (to-csv-record (row-fmt) + (member "name" suppress) + head) (add-cycles data cycles)))))))))) -- 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(-) 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(-) 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 --- Makefile | 4 +- lit/scheduling.org | 3282 ++++++++++++++++++++++++++++++++++++++++++++++ src/hls/RTLBlockgen.v | 2 - src/hls/RTLPargenproof.v | 8 +- 4 files changed, 3287 insertions(+), 9 deletions(-) create mode 100644 lit/scheduling.org diff --git a/Makefile b/Makefile index 5e1ae44..3ff20a6 100644 --- a/Makefile +++ b/Makefile @@ -21,7 +21,7 @@ COQMAKE := $(COQBIN)coq_makefile COQDOCFLAGS := --no-lib-name -l -VS := src/Compiler.v src/Simulator.v src/HLSOpts.v $(foreach d, common hls bourdoncle, src/$(d)/*.v) +VS := src/Compiler.v src/Simulator.v src/HLSOpts.v $(foreach d, common hls bourdoncle, $(wildcard src/$(d)/*.v)) LIT := $(wildcard lit/*.org) PREFIX ?= . @@ -87,7 +87,7 @@ clean:: Makefile.coq rm -f Makefile.coq detangle-all: - $(foreach vs,$(VS),emacs --batch --find-file $(vs) --eval "(progn(require 'org)(require 'ob-tangle)(org-babel-detangle))") + emacs --batch --eval "(progn(require 'org)(require 'ob-tangle)$(foreach vs,$(VS),(org-babel-detangle \"$(vs)\")))" clean:: rm -f */*.v.d */*.glob */*.vo */*~ *~ diff --git a/lit/scheduling.org b/lit/scheduling.org new file mode 100644 index 0000000..7970019 --- /dev/null +++ b/lit/scheduling.org @@ -0,0 +1,3282 @@ +#+title: Scheduling +#+author: Yann Herklotz +#+email: yann@yannherklotz.com + +* Language Definitions + +** RTLBlockInstr +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockInstr.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblockinstr-main +#+begin_src coq +Require Import Coq.micromega.Lia. + +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.Values. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. +Require Import compcert.verilog.Op. + +Require Import Predicate. +Require Import Vericertlib. + +(** * 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]). +*) + +Definition node := positive. + +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. + +(** ** Control-Flow Instruction Definition + +These are the instructions that count as control-flow, and will be placed at the end of the basic +blocks. *) + +Inductive cf_instr : Type := +| RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr +| RBtailcall : signature -> reg + ident -> list reg -> cf_instr +| RBbuiltin : external_function -> list (builtin_arg reg) -> + builtin_res reg -> node -> cf_instr +| RBcond : condition -> list reg -> node -> node -> cf_instr +| RBjumptable : reg -> list node -> cf_instr +| RBreturn : option reg -> cf_instr +| RBgoto : node -> cf_instr +| RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. + +(** ** Helper functions *) + +Fixpoint successors_instr (i : cf_instr) : list node := + match i with + | RBcall sig ros args res s => s :: nil + | RBtailcall sig ros args => nil + | RBbuiltin ef args res s => s :: nil + | RBcond cond args ifso ifnot => ifso :: ifnot :: nil + | RBjumptable arg tbl => tbl + | RBreturn optarg => nil + | RBgoto n => n :: nil + | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil) + end. + +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) + | RBload p chunk addr args dst => + fold_left Pos.max args (Pos.max dst m) + | RBstore p chunk addr args src => + fold_left Pos.max args (Pos.max src m) + | RBsetpred p' c args p => + 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)) + | RBcall sig (inr id) args res s => + fold_left Pos.max args (Pos.max res m) + | RBtailcall sig (inl r) args => + fold_left Pos.max args (Pos.max r m) + | RBtailcall sig (inr id) args => + 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) + | RBcond cond args ifso ifnot => fold_left Pos.max args m + | RBjumptable arg tbl => Pos.max arg m + | RBreturn None => m + | RBreturn (Some arg) => Pos.max arg m + | RBgoto n => m + | RBpred_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. + +(** *** 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; +}. + +(** ** Top-Level Type Definitions *) + +Section DEFINITION. + + Context {bblock_body: Type}. + + 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_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. + +End DEFINITION. + +(** ** Semantics *) + +Section RELSEM. + + Context {bblock_body : Type}. + + Definition genv := Genv.t (@fundef bblock_body) unit. + + 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_RBnop: + 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 + | 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 + | 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 + | 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. + + 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, + 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') + 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) (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') + | 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) + | 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) + | 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) + 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) + | 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'. + +End RELSEM. +#+end_src + +** RTLBlock +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlock.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblock-main +#+begin_src coq +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.RTLBlockInstr. + +Definition bb := list instr. + +Definition bblock := @bblock bb. +Definition code := @code bb. +Definition function := @function bb. +Definition fundef := @fundef bb. +Definition program := @program bb. +Definition funsig := @funsig bb. +Definition stackframe := @stackframe 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). + + 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_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). +#+end_src + +** RTLPar +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPar.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpar-main +#+begin_src coq +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.RTLBlockInstr. + +Definition bb := list (list (list instr)). + +Definition bblock := @bblock bb. +Definition code := @code bb. +Definition function := @function bb. +Definition fundef := @fundef bb. +Definition program := @program bb. +Definition funsig := @funsig bb. +Definition stackframe := @stackframe bb. +Definition state := @state bb. +Definition genv := @genv bb. + +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_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 -> bb -> 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: 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 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') + | 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. +#+end_src +* Scheduler +:PROPERTIES: +:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Schedule.ml +:END: + +#+begin_src ocaml :comments no :padline no :exports none +<> +#+end_src + +#+name: scheduler-main +#+begin_src ocaml +open Printf +open Clflags +open Camlcoq +open Datatypes +open Coqlib +open Maps +open AST +open Kildall +open Op +open RTLBlockInstr +open Predicate +open RTLBlock +open HTL +open Verilog +open HTLgen +open HTLMonad +open HTLMonadExtra + +module SS = Set.Make(P) + +type svtype = + | BBType of int + | VarType of int * int + +type sv = { + sv_type: svtype; + sv_num: int; +} + +let print_sv v = + match v with + | { sv_type = BBType bbi; sv_num = n } -> sprintf "bb%d_%d" bbi n + | { sv_type = VarType (bbi, i); sv_num = n } -> sprintf "var%dn%d_%d" bbi i n + +module G = Graph.Persistent.Digraph.ConcreteLabeled(struct + type t = sv + let compare = compare + let equal = (=) + let hash = Hashtbl.hash +end)(struct + type t = int + let compare = compare + let hash = Hashtbl.hash + let equal = (=) + let default = 0 +end) + +module GDot = Graph.Graphviz.Dot(struct + let graph_attributes _ = [] + let default_vertex_attributes _ = [] + let vertex_name = print_sv + let vertex_attributes _ = [] + let get_subgraph _ = None + let default_edge_attributes _ = [] + let edge_attributes _ = [] + + include G + end) + +module DFG = Graph.Persistent.Digraph.ConcreteLabeled(struct + type t = int * instr + let compare = compare + let equal = (=) + let hash = Hashtbl.hash +end)(struct + type t = int + let compare = compare + let hash = Hashtbl.hash + let equal = (=) + let default = 0 +end) + +module DFGSimp = Graph.Persistent.Graph.Concrete(struct + type t = int * instr + let compare = compare + let equal = (=) + let hash = Hashtbl.hash + end) + +let convert dfg = + DFG.fold_vertex (fun v g -> DFGSimp.add_vertex g v) dfg DFGSimp.empty + |> DFG.fold_edges (fun v1 v2 g -> DFGSimp.add_edge (DFGSimp.add_edge g v1 v2) v2 v1) dfg + +let reg r = sprintf "r%d" (P.to_int r) +let print_pred r = sprintf "p%d" (P.to_int r) + +let print_instr = function + | RBnop -> "" + | RBload (_, _, _, _, r) -> sprintf "load(%s)" (reg r) + | RBstore (_, _, _, _, r) -> sprintf "store(%s)" (reg r) + | RBsetpred (_, _, _, p) -> sprintf "setpred(%s)" (print_pred p) + | RBop (_, op, args, d) -> + (match op, args with + | Omove, _ -> "mov" + | Ointconst n, _ -> sprintf "%s=%ld" (reg d) (camlint_of_coqint n) + | Olongconst n, _ -> sprintf "%s=%LdL" (reg d) (camlint64_of_coqint n) + | Ofloatconst n, _ -> sprintf "%s=%.15F" (reg d) (camlfloat_of_coqfloat n) + | Osingleconst n, _ -> sprintf "%s=%.15Ff" (reg d) (camlfloat_of_coqfloat32 n) + | Oindirectsymbol id, _ -> sprintf "%s=&%s" (reg d) (extern_atom id) + | Ocast8signed, [r1] -> sprintf "%s=int8signed(%s)" (reg d) (reg r1) + | Ocast8unsigned, [r1] -> sprintf "%s=int8unsigned(%s)" (reg d) (reg r1) + | Ocast16signed, [r1] -> sprintf "%s=int16signed(%s)" (reg d) (reg r1) + | Ocast16unsigned, [r1] -> sprintf "%s=int16unsigned(%s)" (reg d) (reg r1) + | Oneg, [r1] -> sprintf "%s=-%s" (reg d) (reg r1) + | Osub, [r1;r2] -> sprintf "%s=%s-%s" (reg d) (reg r1) (reg r2) + | Omul, [r1;r2] -> sprintf "%s=%s*%s" (reg d) (reg r1) (reg r2) + | Omulimm n, [r1] -> sprintf "%s=%s*%ld" (reg d) (reg r1) (camlint_of_coqint n) + | Omulhs, [r1;r2] -> sprintf "%s=mulhs(%s,%s)" (reg d) (reg r1) (reg r2) + | Omulhu, [r1;r2] -> sprintf "%s=mulhu(%s,%s)" (reg d) (reg r1) (reg r2) + | Odiv, [r1;r2] -> sprintf "%s=%s /s %s" (reg d) (reg r1) (reg r2) + | Odivu, [r1;r2] -> sprintf "%s=%s /u %s" (reg d) (reg r1) (reg r2) + | Omod, [r1;r2] -> sprintf "%s=%s %%s %s" (reg d) (reg r1) (reg r2) + | Omodu, [r1;r2] -> sprintf "%s=%s %%u %s" (reg d) (reg r1) (reg r2) + | Oand, [r1;r2] -> sprintf "%s=%s & %s" (reg d) (reg r1) (reg r2) + | Oandimm n, [r1] -> sprintf "%s=%s & %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oor, [r1;r2] -> sprintf "%s=%s | %s" (reg d) (reg r1) (reg r2) + | Oorimm n, [r1] -> sprintf "%s=%s | %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oxor, [r1;r2] -> sprintf "%s=%s ^ %s" (reg d) (reg r1) (reg r2) + | Oxorimm n, [r1] -> sprintf "%s=%s ^ %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Onot, [r1] -> sprintf "%s=not(%s)" (reg d) (reg r1) + | Oshl, [r1;r2] -> sprintf "%s=%s << %s" (reg d) (reg r1) (reg r2) + | Oshlimm n, [r1] -> sprintf "%s=%s << %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshr, [r1;r2] -> sprintf "%s=%s >>s %s" (reg d) (reg r1) (reg r2) + | Oshrimm n, [r1] -> sprintf "%s=%s >>s %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrximm n, [r1] -> sprintf "%s=%s >>x %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshru, [r1;r2] -> sprintf "%s=%s >>u %s" (reg d) (reg r1) (reg r2) + | Oshruimm n, [r1] -> sprintf "%s=%s >>u %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Ororimm n, [r1] -> sprintf "%s=%s ror %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshldimm n, [r1;r2] -> sprintf "%s=(%s, %s) << %ld" (reg d) (reg r1) (reg r2) (camlint_of_coqint n) + | Olea addr, args -> sprintf "%s=addr" (reg d) + | Omakelong, [r1;r2] -> sprintf "%s=makelong(%s,%s)" (reg d) (reg r1) (reg r2) + | Olowlong, [r1] -> sprintf "%s=lowlong(%s)" (reg d) (reg r1) + | Ohighlong, [r1] -> sprintf "%s=highlong(%s)" (reg d) (reg r1) + | Ocast32signed, [r1] -> sprintf "%s=long32signed(%s)" (reg d) (reg r1) + | Ocast32unsigned, [r1] -> sprintf "%s=long32unsigned(%s)" (reg d) (reg r1) + | Onegl, [r1] -> sprintf "%s=-l %s" (reg d) (reg r1) + | Osubl, [r1;r2] -> sprintf "%s=%s -l %s" (reg d) (reg r1) (reg r2) + | Omull, [r1;r2] -> sprintf "%s=%s *l %s" (reg d) (reg r1) (reg r2) + | Omullimm n, [r1] -> sprintf "%s=%s *l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Omullhs, [r1;r2] -> sprintf "%s=mullhs(%s,%s)" (reg d) (reg r1) (reg r2) + | Omullhu, [r1;r2] -> sprintf "%s=mullhu(%s,%s)" (reg d) (reg r1) (reg r2) + | Odivl, [r1;r2] -> sprintf "%s=%s /ls %s" (reg d) (reg r1) (reg r2) + | Odivlu, [r1;r2] -> sprintf "%s=%s /lu %s" (reg d) (reg r1) (reg r2) + | Omodl, [r1;r2] -> sprintf "%s=%s %%ls %s" (reg d) (reg r1) (reg r2) + | Omodlu, [r1;r2] -> sprintf "%s=%s %%lu %s" (reg d) (reg r1) (reg r2) + | Oandl, [r1;r2] -> sprintf "%s=%s &l %s" (reg d) (reg r1) (reg r2) + | Oandlimm n, [r1] -> sprintf "%s=%s &l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Oorl, [r1;r2] -> sprintf "%s=%s |l %s" (reg d) (reg r1) (reg r2) + | Oorlimm n, [r1] -> sprintf "%s=%s |l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Oxorl, [r1;r2] -> sprintf "%s=%s ^l %s" (reg d) (reg r1) (reg r2) + | Oxorlimm n, [r1] -> sprintf "%s=%s ^l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Onotl, [r1] -> sprintf "%s=notl(%s)" (reg d) (reg r1) + | Oshll, [r1;r2] -> sprintf "%s=%s < sprintf "%s=%s < sprintf "%s=%s >>ls %s" (reg d) (reg r1) (reg r2) + | Oshrlimm n, [r1] -> sprintf "%s=%s >>ls %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrxlimm n, [r1] -> sprintf "%s=%s >>lx %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrlu, [r1;r2] -> sprintf "%s=%s >>lu %s" (reg d) (reg r1) (reg r2) + | Oshrluimm n, [r1] -> sprintf "%s=%s >>lu %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Ororlimm n, [r1] -> sprintf "%s=%s rorl %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oleal addr, args -> sprintf "%s=addr" (reg d) + | Onegf, [r1] -> sprintf "%s=negf(%s)" (reg d) (reg r1) + | Oabsf, [r1] -> sprintf "%s=absf(%s)" (reg d) (reg r1) + | Oaddf, [r1;r2] -> sprintf "%s=%s +f %s" (reg d) (reg r1) (reg r2) + | Osubf, [r1;r2] -> sprintf "%s=%s -f %s" (reg d) (reg r1) (reg r2) + | Omulf, [r1;r2] -> sprintf "%s=%s *f %s" (reg d) (reg r1) (reg r2) + | Odivf, [r1;r2] -> sprintf "%s=%s /f %s" (reg d) (reg r1) (reg r2) + | Onegfs, [r1] -> sprintf "%s=negfs(%s)" (reg d) (reg r1) + | Oabsfs, [r1] -> sprintf "%s=absfs(%s)" (reg d) (reg r1) + | Oaddfs, [r1;r2] -> sprintf "%s=%s +fs %s" (reg d) (reg r1) (reg r2) + | Osubfs, [r1;r2] -> sprintf "%s=%s -fs %s" (reg d) (reg r1) (reg r2) + | Omulfs, [r1;r2] -> sprintf "%s=%s *fs %s" (reg d) (reg r1) (reg r2) + | Odivfs, [r1;r2] -> sprintf "%s=%s /fs %s" (reg d) (reg r1) (reg r2) + | Osingleoffloat, [r1] -> sprintf "%s=singleoffloat(%s)" (reg d) (reg r1) + | Ofloatofsingle, [r1] -> sprintf "%s=floatofsingle(%s)" (reg d) (reg r1) + | Ointoffloat, [r1] -> sprintf "%s=intoffloat(%s)" (reg d) (reg r1) + | Ofloatofint, [r1] -> sprintf "%s=floatofint(%s)" (reg d) (reg r1) + | Ointofsingle, [r1] -> sprintf "%s=intofsingle(%s)" (reg d) (reg r1) + | Osingleofint, [r1] -> sprintf "%s=singleofint(%s)" (reg d) (reg r1) + | Olongoffloat, [r1] -> sprintf "%s=longoffloat(%s)" (reg d) (reg r1) + | Ofloatoflong, [r1] -> sprintf "%s=floatoflong(%s)" (reg d) (reg r1) + | Olongofsingle, [r1] -> sprintf "%s=longofsingle(%s)" (reg d) (reg r1) + | Osingleoflong, [r1] -> sprintf "%s=singleoflong(%s)" (reg d) (reg r1) + | Ocmp c, args -> sprintf "%s=cmp" (reg d) + | Osel (c, ty), r1::r2::args -> sprintf "%s=sel" (reg d) + | _, _ -> sprintf "N/a") + +module DFGDot = Graph.Graphviz.Dot(struct + let graph_attributes _ = [] + let default_vertex_attributes _ = [] + let vertex_name = function (i, instr) -> sprintf "\"%d:%s\"" i (print_instr instr) + let vertex_attributes _ = [] + let get_subgraph _ = None + let default_edge_attributes _ = [] + let edge_attributes _ = [] + + include DFG + end) + +module DFGDfs = Graph.Traverse.Dfs(DFG) + +module IMap = Map.Make (struct + type t = int + + let compare = compare +end) + +let gen_vertex instrs i = (i, List.nth instrs i) + +(** The DFG type defines a list of instructions with their data dependencies as [edges], which are + the pairs of integers that represent the index of the instruction in the [nodes]. The edges + always point from left to right. *) + +let print_list f out_chan a = + fprintf out_chan "[ "; + List.iter (fprintf out_chan "%a " f) a; + fprintf out_chan "]" + +let print_tuple out_chan a = + let l, r = a in + fprintf out_chan "(%d,%d)" l r + +(*let print_dfg out_chan dfg = + fprintf out_chan "{ nodes = %a, edges = %a }" + (print_list PrintRTLBlockInstr.print_bblock_body) + dfg.nodes (print_list print_tuple) dfg.edges*) + +let print_dfg = DFGDot.output_graph + +let read_process command = + let buffer_size = 2048 in + let buffer = Buffer.create buffer_size in + let string = Bytes.create buffer_size in + let in_channel = Unix.open_process_in command in + let chars_read = ref 1 in + while !chars_read <> 0 do + chars_read := input in_channel string 0 buffer_size; + Buffer.add_substring buffer (Bytes.to_string string) 0 !chars_read + done; + ignore (Unix.close_process_in in_channel); + Buffer.contents buffer + +let comb_delay = function + | RBnop -> 0 + | RBop (_, op, _, _) -> + (match op with + | Omove -> 0 + | Ointconst _ -> 0 + | Olongconst _ -> 0 + | Ocast8signed -> 0 + | Ocast8unsigned -> 0 + | Ocast16signed -> 0 + | Ocast16unsigned -> 0 + | Oneg -> 0 + | Onot -> 0 + | Oor -> 0 + | Oorimm _ -> 0 + | Oand -> 0 + | Oandimm _ -> 0 + | Oxor -> 0 + | Oxorimm _ -> 0 + | Omul -> 8 + | Omulimm _ -> 8 + | Omulhs -> 8 + | Omulhu -> 8 + | Odiv -> 72 + | Odivu -> 72 + | Omod -> 72 + | Omodu -> 72 + | _ -> 1) + | _ -> 1 + +let pipeline_stages = function + | RBload _ -> 2 + | RBop (_, op, _, _) -> + (match op with + | Odiv -> 32 + | Odivu -> 32 + | Omod -> 32 + | Omodu -> 32 + | _ -> 0) + | _ -> 0 + +(** Add a dependency if it uses a register that was written to previously. *) +let add_dep map i tree dfg curr = + match PTree.get curr tree with + | None -> dfg + | Some ip -> + let ipv = (List.nth map ip) in + 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. + + This function only gathers the RAW constraints, and will therefore only be active for operations + that modify registers, which is this case only affects loads and operations. *) +let accumulate_RAW_deps map dfg curr = + let i, dst_map, graph = dfg in + let acc_dep_instruction rs dst = + ( i + 1, + PTree.set dst i dst_map, + List.fold_left (add_dep map i dst_map) graph rs + ) + in + let acc_dep_instruction_nodst rs = + ( i + 1, + dst_map, + List.fold_left (add_dep map i dst_map) graph rs) + in + 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) + +(** Finds the next write to the [dst] register. This is a small optimisation so that only one + dependency is generated for a data dependency. *) +let rec find_next_dst_write i dst i' curr = + let check_dst dst' curr' = + if dst = dst' then Some (i, i') + else find_next_dst_write i dst (i' + 1) curr' + in + match curr with + | [] -> None + | RBop (_, _, _, dst') :: curr' -> check_dst dst' curr' + | RBload (_, _, _, _, dst') :: curr' -> check_dst dst' curr' + | _ :: curr' -> find_next_dst_write i dst (i' + 1) curr' + +let rec find_all_next_dst_read i dst i' curr = + let check_dst rs curr' = + if List.exists (fun x -> x = dst) rs + then (i, i') :: find_all_next_dst_read i dst (i' + 1) curr' + else find_all_next_dst_read i dst (i' + 1) curr' + in + match curr with + | [] -> [] + | RBop (_, _, rs, _) :: curr' -> check_dst rs curr' + | RBload (_, _, _, rs, _) :: curr' -> check_dst rs curr' + | RBstore (_, _, _, rs, src) :: curr' -> check_dst (src :: rs) curr' + | RBnop :: curr' -> find_all_next_dst_read i dst (i' + 1) curr' + | RBsetpred (_, _, rs, _) :: curr' -> check_dst rs curr' + +let drop i lst = + let rec drop' i' lst' = + match lst' with + | _ :: ls -> if i' = i then ls else drop' (i' + 1) ls + | [] -> [] + in + if i = 0 then lst else drop' 1 lst + +let take i lst = + let rec take' i' lst' = + match lst' with + | l :: ls -> if i' = i then [ l ] else l :: take' (i' + 1) ls + | [] -> [] + in + if i = 0 then [] else take' 1 lst + +let rec next_store i = function + | [] -> None + | RBstore (_, _, _, _, _) :: _ -> Some i + | _ :: rst -> next_store (i + 1) rst + +let rec next_load i = function + | [] -> None + | RBload (_, _, _, _, _) :: _ -> Some i + | _ :: rst -> next_load (i + 1) rst + +let accumulate_RAW_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBload (_, _, _, _, _) -> ( + 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) ) + | _ -> dfg + +let accumulate_WAR_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBstore (_, _, _, _, _) -> ( + match next_load 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) ) + | _ -> dfg + +let accumulate_WAW_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | 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)) + | _ -> dfg + +(** Predicate dependencies. *) + +let rec in_predicate p p' = + match p' with + | Plit p'' -> P.to_int p = P.to_int (snd p'') + | Pand (p1, p2) -> in_predicate p p1 || in_predicate p p2 + | Por (p1, p2) -> in_predicate p p1 || in_predicate p p2 + | Ptrue -> false + | Pfalse -> false + +let get_predicate = function + | RBop (p, _, _, _) -> p + | RBload (p, _, _, _, _) -> p + | RBstore (p, _, _, _, _) -> p + | RBsetpred (p, _, _, _) -> p + | _ -> None + +let rec next_setpred p i = function + | [] -> None + | RBsetpred (_, _, _, p') :: rst -> + if in_predicate p' p then + Some i + else + next_setpred p (i + 1) rst + | _ :: rst -> next_setpred p (i + 1) rst + +let rec next_preduse p i instr= + let next p' rst = + if in_predicate p p' then + Some i + else + next_preduse p (i + 1) rst + in + match instr with + | [] -> None + | 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 = + let i, curr = curri in + match get_predicate curr with + | Some p -> ( + match next_setpred p 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) ) + | _ -> dfg + +let accumulate_WAR_pred_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBsetpred (_, _, _, p) -> ( + match next_preduse p 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) ) + | _ -> dfg + +let accumulate_WAW_pred_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBsetpred (_, _, _, p) -> ( + match next_setpred (Plit (true, p)) 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) ) + | _ -> dfg + +(** This function calculates the WAW dependencies, which happen when two writes are ordered one + after another and therefore have to be kept in that order. This accumulation might be redundant + if register renaming is done before hand, because then these dependencies can be avoided. *) +let accumulate_WAW_deps instrs dfg curri = + let i, curr = curri in + let dst_dep dst = + match find_next_dst_write i dst (i + 1) (drop (i + 1) instrs) with + | Some (a, b) -> DFG.add_edge dfg (gen_vertex instrs a) (gen_vertex instrs b) + | _ -> dfg + in + match curr with + | RBop (_, _, _, dst) -> dst_dep dst + | RBload (_, _, _, _, dst) -> dst_dep dst + | RBstore (_, _, _, _, _) -> ( + match next_store (i + 1) (drop (i + 1) instrs) with + | None -> dfg + | Some i' -> DFG.add_edge dfg (gen_vertex instrs i) (gen_vertex instrs i') ) + | _ -> dfg + +let accumulate_WAR_deps instrs dfg curri = + let i, curr = curri in + let dst_dep dst = + let dep_list = find_all_next_dst_read i dst 0 (take i instrs |> List.rev) + |> List.map (function (d, d') -> (i - d' - 1, d)) + in + List.fold_left (fun g -> + function (d, d') -> DFG.add_edge g (gen_vertex instrs d) (gen_vertex instrs d')) dfg dep_list + in + match curr with + | RBop (_, _, _, dst) -> dst_dep dst + | RBload (_, _, _, _, dst) -> dst_dep dst + | _ -> dfg + +let assigned_vars vars = function + | RBnop -> vars + | RBop (_, _, _, dst) -> dst :: vars + | RBload (_, _, _, _, dst) -> dst :: vars + | RBstore (_, _, _, _, _) -> vars + | RBsetpred (_, _, _, _) -> vars + +let get_pred = function + | RBnop -> None + | RBop (op, _, _, _) -> op + | RBload (op, _, _, _, _) -> op + | RBstore (op, _, _, _, _) -> op + | RBsetpred (_, _, _, _) -> None + +let independant_pred p p' = + match sat_pred_simple (Pand (p, p')) with + | None -> true + | _ -> false + +let check_dependent op1 op2 = + match op1, op2 with + | Some p, Some p' -> not (independant_pred p p') + | _, _ -> true + +let remove_unnecessary_deps graph = + let is_dependent v1 v2 g = + let (_, instr1) = v1 in + let (_, instr2) = v2 in + if check_dependent (get_pred instr1) (get_pred instr2) + then g + else DFG.remove_edge g v1 v2 + in + DFG.fold_edges is_dependent graph graph + +(** All the nodes in the DFG have to come after the source of the basic block, and should terminate + before the sink of the basic block. After that, there should be constraints for data + dependencies between nodes. *) +let gather_bb_constraints debug bb = + let ibody = List.mapi (fun i a -> (i, a)) bb.bb_body in + let dfg = List.fold_left (fun dfg v -> DFG.add_vertex dfg v) DFG.empty ibody in + let _, _, dfg' = + List.fold_left (accumulate_RAW_deps ibody) + (0, PTree.empty, dfg) + bb.bb_body + in + let dfg'' = List.fold_left (fun dfg f -> List.fold_left (f bb.bb_body) dfg ibody) dfg' + [ accumulate_WAW_deps; + accumulate_WAR_deps; + accumulate_RAW_mem_deps; + accumulate_WAR_mem_deps; + accumulate_WAW_mem_deps; + accumulate_RAW_pred_deps; + accumulate_WAR_pred_deps; + accumulate_WAW_pred_deps + ] + in + let dfg''' = remove_unnecessary_deps dfg'' in + (List.length bb.bb_body, dfg''', successors_instr bb.bb_exit) + +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 + then G.add_edge_e g (encode_bb n 0, 0, encode_var n (fst v1) 0) + else g) |> + (fun g' -> + if DFG.out_degree dfg v1 = 0 + 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, 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 = + List.fold_left (fun g n' -> + G.add_edge_e g (encode_bb n 1, -1, encode_bb n' 0) + ) constr succs + +module BFDFG = Graph.Path.BellmanFord(DFG)(struct + include DFG + type t = int + let weight = DFG.E.label + let compare = compare + let add = (+) + let zero = 0 + end) + +module TopoDFG = Graph.Topological.Make(DFG) + +let negate_graph constr = + DFG.fold_edges_e (function (v1, e, v2) -> fun g -> + DFG.add_edge_e g (v1, -e, v2) + ) constr DFG.empty + +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 + |> 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 maxf))) - 1), + encode_var n (fst v') 0) + ) constr (DFG.fold_vertex (fun v l -> + 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 = + | Mem + | SDiv + | UDiv + +type resources = { + res_mem: DFG.V.t list; + res_udiv: DFG.V.t list; + res_sdiv: DFG.V.t list; +} + +let find_resource = function + | RBload _ -> Some Mem + | RBstore _ -> Some Mem + | RBop (_, op, _, _) -> + ( match op with + | Odiv -> Some SDiv + | Odivu -> Some UDiv + | Omod -> Some SDiv + | Omodu -> Some UDiv + | _ -> None ) + | _ -> None + +let add_resource_constr n dfg constr = + let res = TopoDFG.fold (function (i, instr) -> + function {res_mem = ml; res_sdiv = sdl; res_udiv = udl} as r -> + match find_resource instr with + | Some SDiv -> {r with res_sdiv = (i, instr) :: sdl} + | Some UDiv -> {r with res_udiv = (i, instr) :: udl} + | Some Mem -> {r with res_mem = (i, instr) :: ml} + | None -> r + ) dfg {res_mem = []; res_sdiv = []; res_udiv = []} + in + let get_constraints l g = List.fold_left (fun gv v' -> + match gv with + | (g, None) -> (g, Some v') + | (g, Some v) -> + (G.add_edge_e g (encode_var n (fst v) 0, -1, encode_var n (fst v') 0), Some v') + ) (g, None) l |> fst + in + get_constraints (List.rev res.res_mem) constr + |> get_constraints (List.rev res.res_udiv) + |> get_constraints (List.rev res.res_sdiv) + +let gather_cfg_constraints c constr curr = + let (n, dfg) = curr in + match PTree.get (P.of_int n) c with + | 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 + |> List.filter (fun n' -> n' < n)) + |> add_cycle_constr 8 n dfg + |> add_resource_constr n dfg + +let rec intersperse s = function + | [] -> [] + | [ a ] -> [ a ] + | x :: xs -> x :: s :: intersperse s xs + +let print_objective constr = + let vars = G.fold_vertex (fun v1 l -> + match v1 with + | { sv_type = VarType _; sv_num = 0 } -> print_sv v1 :: l + | _ -> l + ) constr [] + in + "min: " ^ String.concat "" (intersperse " + " vars) ^ ";\n" + +let print_lp constr = + print_objective constr ^ + (G.fold_edges_e (function (e1, w, e2) -> fun s -> + s ^ sprintf "%s - %s <= %d;\n" (print_sv e1) (print_sv e2) w + ) constr "" |> + G.fold_vertex (fun v1 s -> + s ^ sprintf "int %s;\n" (print_sv v1) + ) constr) + +let update_schedule v = function Some l -> Some (v :: l) | None -> Some [ v ] + +let parse_soln (tree, bbtree) s = + let r = Str.regexp "var\\([0-9]+\\)n\\([0-9]+\\)_0[ ]+\\([0-9]+\\)" in + 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 (fn, oc) = Filename.open_temp_file "vericert_" "_lp_solve" in + fprintf oc "%s\n" (print_lp constr); + close_out oc; + + let res = Str.split (Str.regexp_string "\n") (read_process ("lp_solve " ^ fn)) + |> drop 3 + |> List.fold_left parse_soln (IMap.empty, 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 + List.fold_left (fun g v -> + List.fold_left (fun g' v' -> + let edges = DFG.find_all_edges dfg v v' in + List.fold_left DFG.add_edge_e g' edges + ) g l + ) dfg' l + +let rec all_successors dfg v = + List.concat (List.fold_left (fun l v -> + all_successors dfg v :: l + ) [] (DFG.succ dfg v)) + +let order_instr dfg = + DFG.fold_vertex (fun v li -> + if DFG.in_degree dfg v = 0 + then (List.map snd (v :: all_successors dfg v)) :: li + else li + ) dfg [] + +let combine_bb_schedule schedule s = + let i, st = s in + IMap.update st (update_schedule i) schedule + +(**let add_el dfg i l = + List.*) + +let check_in el = + List.exists (List.exists ((=) el)) + +let all_dfs dfg = + let roots = DFG.fold_vertex (fun v li -> + if DFG.in_degree dfg v = 0 then v :: li else li + ) dfg [] in + let dfg' = DFG.fold_edges (fun v1 v2 g -> DFG.add_edge g v2 v1) dfg dfg in + List.fold_left (fun a el -> + if check_in el a then a else + (DFGDfs.fold_component (fun v l -> v :: l) [] dfg' el) :: a) [] roots + +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 = + 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 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 + |> 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 -> + all_dfs x + |> List.map (subgraph x) + |> List.map (fun y -> + TopoDFG.fold (fun i l -> snd i :: l) y [] + |> List.rev))) body2 + (*|> (fun x -> TopoDFG.fold (fun i l -> snd i :: l) x []) + |> List.rev) body2*) + in + { bb_body = final_body2; + bb_exit = ctrl_flow + } + in + PTree.map f c + +let schedule entry (c : RTLBlock.bb RTLBlockInstr.code) = + let debug = true in + let transf_graph (_, dfg, _) = dfg in + let c' = PTree.map1 (fun x -> gather_bb_constraints false x |> transf_graph) c in + (*let _ = if debug then PTree.map (fun r o -> printf "##### %d #####\n%a\n\n" (P.to_int r) print_dfg o) c' else PTree.empty in*) + let cgraph = PTree.elements c' + |> List.map (function (x, y) -> (P.to_int x, y)) + |> List.fold_left (gather_cfg_constraints c) G.empty + in + let graph = open_out "constr_graph.dot" in + fprintf graph "%a\n" GDot.output_graph cgraph; + close_out graph; + let schedule' = solve_constraints cgraph in + (**IMap.iter (fun a b -> printf "##### %d #####\n%a\n\n" a (print_list print_tuple) b) schedule';*) + (**printf "Schedule: %a\n" (fun a x -> IMap.iter (fun d -> fprintf a "%d: %a\n" d (print_list print_tuple)) x) schedule';*) + transf_rtlpar c c' schedule' + +let rec find_reachable_states c e = + match PTree.get e c with + | Some { bb_exit = ex; _ } -> + e :: List.fold_left (fun x a -> List.concat [x; find_reachable_states c a]) [] + (successors_instr ex |> List.filter (fun x -> P.lt x e)) + | None -> assert false + +let add_to_tree c nt i = + match PTree.get i c with + | Some p -> PTree.set i p nt + | None -> assert false + +let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = + let scheduled = schedule f.fn_entrypoint f.fn_code in + let reachable = find_reachable_states scheduled f.fn_entrypoint + |> List.to_seq |> SS.of_seq |> SS.to_seq |> List.of_seq in + { fn_sig = f.fn_sig; + fn_params = f.fn_params; + fn_stacksize = f.fn_stacksize; + fn_code = scheduled (*List.fold_left (add_to_tree scheduled) PTree.empty reachable*); + fn_entrypoint = f.fn_entrypoint + } +#+end_src + +* RTLBlockgen +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgen.v +:END: + +Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]]. + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblockgen-main +#+begin_src coq +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. + +(** [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.*) + +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. + +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. +#+end_src + +* RTLBlockgenproof +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +** Imports + +#+name: rtlblockgenproof-imports +#+begin_src coq +Require compcert.backend.RTL. +Require Import compcert.common.AST. +Require Import compcert.lib.Maps. + +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLBlockgen. +#+end_src + +** Match states + +The ~match_states~ predicate describes which states are equivalent between the two languages, in this +case ~RTL~ and ~RTLBlock~. + +#+name: rtlblockgenproof-match-states +#+begin_src coq +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). +#+end_src + +** Correctness + +#+name: rtlblockgenproof-correctness +#+begin_src coq +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. +#+end_src +* RTLPargen +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargen.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpargen-main +#+begin_src coq +Require Import compcert.backend.Registers. +Require Import compcert.common.AST. +Require Import compcert.common.Globalenvs. +Require Import compcert.common.Memory. +Require Import compcert.common.Values. +Require Import compcert.lib.Floats. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. +Require compcert.verilog.Op. + +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLPar. +Require Import vericert.hls.RTLBlockInstr. +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. + +(** ** 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) + end. + +Fixpoint replicate {A} (n: nat) (l: A) := + match n with + | O => nil + | S n => l :: replicate n l + end. + +Definition merge''' x y := + match x, y with + | Some p1, Some p2 => Some (Pand p1 p2) + | Some p, None | None, Some p => Some p + | None, None => None + end. + +Definition merge'' x := + match x with + | ((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 := + match pa, pf with + | (p, a), (p', f) => (merge''' p p', f a) + end. + +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. + +(*map (fun x => (fst x, Econs (snd x) Enil)) pel*) +Definition merge' (pel: pred_expr) (tpel: predicated expression_list) := + predicated_map (uncurry Econs) (predicated_prod pel tpel). + +Fixpoint merge (pel: list pred_expr): predicated expression_list := + match pel with + | nil => NE.singleton (T, Enil) + | a :: b => merge' a (merge b) + end. + +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 := + 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 := + 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 := + NE.map (fun x => (fst x, (snd x) pa pb pc)) pf. + +Definition predicated_from_opt {A: Type} (p: option pred_op) (a: A) := + match p with + | Some p' => NE.singleton (p', a) + | None => NE.singleton (T, a) + end. + +#[local] Open Scope non_empty_scope. +#[local] Open Scope pred_op. + +Fixpoint NEfold_left {A B} (f: A -> B -> A) (l: NE.non_empty B) (a: A) : A := + match l with + | NE.singleton a' => f a a' + | a' ::| b => NEfold_left f b (f a a') + end. + +Fixpoint NEapp {A} (l m: NE.non_empty A) := + match l with + | NE.singleton a => a ::| m + | a ::| b => a ::| NEapp b m + end. + +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. + +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 + +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 + | RBop p op rl r => + f # (Reg r) <- + (app_predicated p + (f # (Reg r)) + (map_predicated (pred_ret (Eop op)) (merge (list_translation rl f)))) + | RBload p chunk addr rl r => + f # (Reg r) <- + (app_predicated p + (f # (Reg r)) + (map_predicated + (map_predicated (pred_ret (Eload chunk addr)) (merge (list_translation rl f))) + (f # Mem))) + | RBstore p chunk addr rl r => + f # Mem <- + (app_predicated p + (f # Mem) + (map_predicated + (map_predicated + (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)))) + 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. + +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. + +(** 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. *) + +Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := + match bb with + | nil => + match bbt with + | nil => true + | _ => false + end + | _ => true + end. + +Definition schedule_oracle (bb: RTLBlock.bblock) (bbt: RTLPar.bblock) : bool := + check (abstract_sequence empty (bb_body bb)) + (abstract_sequence empty (concat (concat (bb_body bbt)))) && + check_control_flow_instr (bb_exit bb) (bb_exit bbt) && + empty_trees (bb_body bb) (bb_body bbt). + +Definition check_scheduled_trees := beq2 schedule_oracle. + +Ltac solve_scheduled_trees_correct := + intros; unfold check_scheduled_trees in *; + match goal with + | [ H: context[beq2 _ _ _], x: positive |- _ ] => + rewrite beq2_correct in H; specialize (H x) + end; repeat destruct_match; crush. + +Lemma check_scheduled_trees_correct: + forall f1 f2 x y1, + check_scheduled_trees f1 f2 = true -> + PTree.get x f1 = Some y1 -> + exists y2, PTree.get x f2 = Some y2 /\ schedule_oracle y1 y2 = true. +Proof. solve_scheduled_trees_correct; eexists; crush. Qed. + +Lemma check_scheduled_trees_correct2: + forall f1 f2 x, + check_scheduled_trees f1 f2 = true -> + PTree.get x f1 = None -> + PTree.get x f2 = None. +Proof. solve_scheduled_trees_correct. Qed. + +(** ** Top-level Functions *) + +Parameter schedule : RTLBlock.function -> 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) + f.(fn_params) + f.(fn_stacksize) + tfcode + f.(fn_entrypoint)) + else + 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. +#+end_src +* RTLPargenproof +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargenproof.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpargenproof-imports +#+begin_src coq +Require Import compcert.backend.Registers. +Require Import compcert.common.AST. +Require Import compcert.common.Errors. +Require Import compcert.common.Linking. +Require Import compcert.common.Globalenvs. +Require Import compcert.common.Memory. +Require Import compcert.common.Values. +Require Import compcert.lib.Maps. + +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLPar. +Require Import vericert.hls.RTLBlockInstr. +Require Import vericert.hls.RTLPargen. +Require Import vericert.hls.Predicate. +Require Import vericert.hls.Abstr. + +#[local] Open Scope positive. +#[local] Open Scope forest. +#[local] Open Scope pred_op. +#+end_src + +** RTLBlock to abstract translation + +Correctness of translation from RTLBlock to the abstract interpretation language. + +#+name: rtlpargenproof-main +#+begin_src coq +(*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. + +Inductive state_lessdef : instr_state -> instr_state -> Prop := + state_lessdef_intro : + forall rs1 rs2 m1, + (forall x, rs1 !! x = rs2 !! x) -> + state_lessdef (mk_instr_state rs1 m1) (mk_instr_state rs2 m1). + +Ltac inv_simp := + repeat match goal with + | H: exists _, _ |- _ => inv H + end; simplify. + +*) + +Definition check_dest i r' := + match i with + | RBop p op rl r => (r =? r')%positive + | RBload p chunk addr rl r => (r =? r')%positive + | _ => false + end. + +Lemma check_dest_dec i r : {check_dest i r = true} + {check_dest i r = false}. +Proof. destruct (check_dest i r); tauto. Qed. + +Fixpoint check_dest_l l r := + match l with + | nil => false + | a :: b => check_dest a r || check_dest_l b r + end. + +Lemma check_dest_l_forall : + forall l r, + check_dest_l l r = false -> + Forall (fun x => check_dest x r = false) l. +Proof. induction l; crush. Qed. + +Lemma check_dest_l_dec i r : {check_dest_l i r = true} + {check_dest_l i r = false}. +Proof. destruct (check_dest_l i r); tauto. Qed. + +Lemma check_dest_update : + forall f i r, + check_dest i r = false -> + (update f i) # (Reg r) = f # (Reg r). +Proof. + destruct i; crush; try apply Pos.eqb_neq in H; apply genmap1; crush. +Qed. + +Lemma check_dest_l_forall2 : + forall l r, + Forall (fun x => check_dest x r = false) l -> + check_dest_l l r = false. +Proof. + induction l; crush. + inv H. apply orb_false_intro; crush. +Qed. + +Lemma check_dest_l_ex2 : + forall l r, + (exists a, In a l /\ check_dest a r = true) -> + check_dest_l l r = true. +Proof. + induction l; crush. + specialize (IHl r). inv H. + apply orb_true_intro; crush. + apply orb_true_intro; crush. + right. apply IHl. exists x. auto. +Qed. + +Lemma check_list_l_false : + forall l x r, + check_dest_l (l ++ x :: nil) r = false -> + check_dest_l l r = false /\ check_dest x r = false. +Proof. + simplify. + apply check_dest_l_forall in H. apply Forall_app in H. + simplify. apply check_dest_l_forall2; auto. + apply check_dest_l_forall in H. apply Forall_app in H. + simplify. inv H1. auto. +Qed. + +Lemma check_dest_l_ex : + forall l r, + check_dest_l l r = true -> + exists a, In a l /\ check_dest a r = true. +Proof. + induction l; crush. + destruct (check_dest a r) eqn:?; try solve [econstructor; crush]. + simplify. + exploit IHl. apply H. simplify. econstructor. simplify. right. eassumption. + auto. +Qed. + +Lemma check_list_l_true : + forall l x r, + check_dest_l (l ++ x :: nil) r = true -> + check_dest_l l r = true \/ check_dest x r = true. +Proof. + simplify. + apply check_dest_l_ex in H; simplify. + apply in_app_or in H. inv H. left. + apply check_dest_l_ex2. exists x0. auto. + inv H0; auto. +Qed. + +Lemma check_dest_l_dec2 l r : + {Forall (fun x => check_dest x r = false) l} + + {exists a, In a l /\ check_dest a r = true}. +Proof. + destruct (check_dest_l_dec l r); [right | left]; + auto using check_dest_l_ex, check_dest_l_forall. +Qed. + +Lemma abstr_comp : + forall l i f x x0, + abstract_sequence f (l ++ i :: nil) = x -> + abstract_sequence f l = x0 -> + x = update x0 i. +Proof. induction l; intros; crush; eapply IHl; eauto. Qed. + +(* + +Lemma gen_list_base: + forall FF ge sp l rs exps st1, + (forall x, @sem_value FF ge sp st1 (exps # (Reg x)) (rs !! x)) -> + sem_val_list ge sp st1 (list_translation l exps) rs ## l. +Proof. + induction l. + intros. simpl. constructor. + intros. simpl. eapply Scons; eauto. +Qed. + +Lemma check_dest_update2 : + forall f r rl op p, + (update f (RBop p op rl r)) # (Reg r) = Eop op (list_translation rl f) (f # Mem). +Proof. crush; rewrite map2; auto. Qed. + +Lemma check_dest_update3 : + forall f r rl p addr chunk, + (update f (RBload p chunk addr rl r)) # (Reg r) = Eload chunk addr (list_translation rl f) (f # Mem). +Proof. crush; rewrite map2; auto. Qed. + +Lemma abstract_seq_correct_aux: + forall FF ge sp i st1 st2 st3 f, + @step_instr FF ge sp st3 i st2 -> + sem ge sp st1 f st3 -> + sem ge sp st1 (update f i) st2. +Proof. + intros; inv H; simplify. + { simplify; eauto. } (*apply match_states_refl. }*) + { inv H0. inv H6. destruct st1. econstructor. simplify. + constructor. intros. + destruct (resource_eq (Reg res) (Reg x)). inv e. + rewrite map2. econstructor. eassumption. apply gen_list_base; eauto. + rewrite Regmap.gss. eauto. + assert (res <> x). { unfold not in *. intros. apply n. rewrite H0. auto. } + rewrite Regmap.gso by auto. + rewrite genmap1 by auto. auto. + + rewrite genmap1; crush. } + { inv H0. inv H7. constructor. constructor. intros. + destruct (Pos.eq_dec dst x); subst. + rewrite map2. econstructor; eauto. + apply gen_list_base. auto. rewrite Regmap.gss. auto. + rewrite genmap1. rewrite Regmap.gso by auto. auto. + unfold not in *; intros. inv H0. auto. + rewrite genmap1; crush. + } + { inv H0. inv H7. constructor. constructor; intros. + rewrite genmap1; crush. + rewrite map2. econstructor; eauto. + apply gen_list_base; auto. + } +Qed. + +Lemma regmap_list_equiv : + forall A (rs1: Regmap.t A) rs2, + (forall x, rs1 !! x = rs2 !! x) -> + forall rl, rs1##rl = rs2##rl. +Proof. induction rl; crush. Qed. + +Lemma sem_update_Op : + forall A ge sp st f st' r l o0 o m rs v, + @sem A ge sp st f st' -> + Op.eval_operation ge sp o0 rs ## l m = Some v -> + match_states st' (mk_instr_state rs m) -> + exists tst, + sem ge sp st (update f (RBop o o0 l r)) tst /\ match_states (mk_instr_state (Regmap.set r v rs) m) tst. +Proof. + intros. inv H1. simplify. + destruct st. + econstructor. simplify. + { constructor. + { constructor. intros. destruct (Pos.eq_dec x r); subst. + { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. + { inv H9. eapply gen_list_base; eauto. } + { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } + { inv H. rewrite genmap1; crush. eauto. } } + { constructor; eauto. intros. + destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed. + +Lemma sem_update_load : + forall A ge sp st f st' r o m a l m0 rs v a0, + @sem A ge sp st f st' -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.loadv m m0 a0 = Some v -> + match_states st' (mk_instr_state rs m0) -> + exists tst : instr_state, + sem ge sp st (update f (RBload o m a l r)) tst + /\ match_states (mk_instr_state (Regmap.set r v rs) m0) tst. +Proof. + intros. inv H2. pose proof H. inv H. inv H9. + destruct st. + econstructor; simplify. + { constructor. + { constructor. intros. + destruct (Pos.eq_dec x r); subst. + { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. + rewrite <- H6. eauto. + instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. auto. } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. } } + { rewrite genmap1; crush. eauto. } } + { constructor; auto; intros. destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed. + +Lemma sem_update_store : + forall A ge sp a0 m a l r o f st m' rs m0 st', + @sem A ge sp st f st' -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.storev m m0 a0 rs !! r = Some m' -> + match_states st' (mk_instr_state rs m0) -> + exists tst, sem ge sp st (update f (RBstore o m a l r)) tst + /\ match_states (mk_instr_state rs m') tst. +Proof. + intros. inv H2. pose proof H. inv H. inv H9. + destruct st. + econstructor; simplify. + { econstructor. + { econstructor; intros. rewrite genmap1; crush. } + { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. rewrite <- H6. + eauto. specialize (H6 r). rewrite H6. eauto. } } + { econstructor; eauto. } +Qed. + +Lemma sem_update : + forall A ge sp st x st' st'' st''' f, + sem ge sp st f st' -> + match_states st' st''' -> + @step_instr A ge sp st''' x st'' -> + exists tst, sem ge sp st (update f x) tst /\ match_states st'' tst. +Proof. + intros. destruct x; inv H1. + { econstructor. split. + apply sem_update_RBnop. eassumption. + apply match_states_commut. auto. } + { eapply sem_update_Op; eauto. } + { eapply sem_update_load; eauto. } + { eapply sem_update_store; eauto. } +Qed. + +Lemma sem_update2_Op : + forall A ge sp st f r l o0 o m rs v, + @sem A ge sp st f (mk_instr_state rs m) -> + Op.eval_operation ge sp o0 rs ## l m = Some v -> + sem ge sp st (update f (RBop o o0 l r)) (mk_instr_state (Regmap.set r v rs) m). +Proof. + intros. destruct st. constructor. + inv H. inv H6. + { constructor; intros. simplify. + destruct (Pos.eq_dec r x); subst. + { rewrite map2. econstructor. eauto. + apply gen_list_base. eauto. + rewrite Regmap.gss. auto. } + { rewrite genmap1; crush. rewrite Regmap.gso; auto. } } + { simplify. rewrite genmap1; crush. inv H. eauto. } +Qed. + +Lemma sem_update2_load : + forall A ge sp st f r o m a l m0 rs v a0, + @sem A ge sp st f (mk_instr_state rs m0) -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.loadv m m0 a0 = Some v -> + sem ge sp st (update f (RBload o m a l r)) (mk_instr_state (Regmap.set r v rs) m0). +Proof. + intros. simplify. inv H. inv H7. constructor. + { constructor; intros. destruct (Pos.eq_dec r x); subst. + { rewrite map2. rewrite Regmap.gss. econstructor; eauto. + apply gen_list_base; eauto. } + { rewrite genmap1; crush. rewrite Regmap.gso; eauto. } + } + { simplify. rewrite genmap1; crush. } +Qed. + +Lemma sem_update2_store : + forall A ge sp a0 m a l r o f st m' rs m0, + @sem A ge sp st f (mk_instr_state rs m0) -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.storev m m0 a0 rs !! r = Some m' -> + sem ge sp st (update f (RBstore o m a l r)) (mk_instr_state rs m'). +Proof. + intros. simplify. inv H. inv H7. constructor; simplify. + { econstructor; intros. rewrite genmap1; crush. } + { rewrite map2. econstructor; eauto. apply gen_list_base; eauto. } +Qed. + +Lemma sem_update2 : + forall A ge sp st x st' st'' f, + sem ge sp st f st' -> + @step_instr A ge sp st' x st'' -> + sem ge sp st (update f x) st''. +Proof. + intros. + destruct x; inv H0; + eauto using sem_update_RBnop, sem_update2_Op, sem_update2_load, sem_update2_store. +Qed. + +Lemma abstr_sem_val_mem : + forall A B ge tge st tst sp a, + ge_preserved ge tge -> + forall v m, + (@sem_mem A ge sp st a m /\ match_states st tst -> @sem_mem B tge sp tst a m) /\ + (@sem_value A ge sp st a v /\ match_states st tst -> @sem_value B tge sp tst a v). +Proof. + intros * H. + apply expression_ind2 with + + (P := fun (e1: expression) => + forall v m, + (@sem_mem A ge sp st e1 m /\ match_states st tst -> @sem_mem B tge sp tst e1 m) /\ + (@sem_value A ge sp st e1 v /\ match_states st tst -> @sem_value B tge sp tst e1 v)) + + (P0 := fun (e1: expression_list) => + forall lv, @sem_val_list A ge sp st e1 lv /\ match_states st tst -> @sem_val_list B tge sp tst e1 lv); + simplify; intros; simplify. + { inv H1. inv H2. constructor. } + { inv H2. inv H1. rewrite H0. constructor. } + { inv H3. } + { inv H3. inv H4. econstructor. apply H1; auto. simplify. eauto. constructor. auto. auto. + apply H0; simplify; eauto. constructor; eauto. + unfold ge_preserved in *. simplify. rewrite <- H2. auto. + } + { inv H3. } + { inv H3. inv H4. econstructor. apply H1; eauto; simplify; eauto. constructor; eauto. + apply H0; simplify; eauto. constructor; eauto. + inv H. rewrite <- H4. eauto. + auto. + } + { inv H4. inv H5. econstructor. apply H0; eauto. simplify; eauto. constructor; eauto. + apply H2; eauto. simplify; eauto. constructor; eauto. + apply H1; eauto. simplify; eauto. constructor; eauto. + inv H. rewrite <- H5. eauto. auto. + } + { inv H4. } + { inv H1. constructor. } + { inv H3. constructor; auto. apply H0; eauto. apply Mem.empty. } +Qed. + +Lemma abstr_sem_value : + forall a A B ge tge sp st tst v, + @sem_value A ge sp st a v -> + ge_preserved ge tge -> + match_states st tst -> + @sem_value B tge sp tst a v. +Proof. intros; eapply abstr_sem_val_mem; eauto; apply Mem.empty. Qed. + +Lemma abstr_sem_mem : + forall a A B ge tge sp st tst v, + @sem_mem A ge sp st a v -> + ge_preserved ge tge -> + match_states st tst -> + @sem_mem B tge sp tst a v. +Proof. intros; eapply abstr_sem_val_mem; eauto. Qed. + +Lemma abstr_sem_regset : + forall a a' A B ge tge sp st tst rs, + @sem_regset A ge sp st a rs -> + ge_preserved ge tge -> + (forall x, a # x = a' # x) -> + match_states st tst -> + exists rs', @sem_regset B tge sp tst a' rs' /\ (forall x, rs !! x = rs' !! x). +Proof. + inversion 1; intros. + inv H7. + econstructor. simplify. econstructor. intros. + eapply abstr_sem_value; eauto. rewrite <- H6. + eapply H0. constructor; eauto. + auto. +Qed. + +Lemma abstr_sem : + forall a a' A B ge tge sp st tst st', + @sem A ge sp st a st' -> + ge_preserved ge tge -> + (forall x, a # x = a' # x) -> + match_states st tst -> + exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. +Proof. + inversion 1; subst; intros. + inversion H4; subst. + exploit abstr_sem_regset; eauto; inv_simp. + do 3 econstructor; eauto. + rewrite <- H3. + eapply abstr_sem_mem; eauto. +Qed. + +Lemma abstract_execution_correct': + forall A B ge tge sp st' a a' st tst, + @sem A ge sp st a st' -> + ge_preserved ge tge -> + check a a' = true -> + match_states st tst -> + exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. +Proof. + intros; + pose proof (check_correct a a' H1); + eapply abstr_sem; eauto. +Qed. + +Lemma states_match : + forall st1 st2 st3 st4, + match_states st1 st2 -> + match_states st2 st3 -> + match_states st3 st4 -> + match_states st1 st4. +Proof. + intros * H1 H2 H3; destruct st1; destruct st2; destruct st3; destruct st4. + inv H1. inv H2. inv H3; constructor. + unfold regs_lessdef in *. intros. + repeat match goal with + | H: forall _, _, r : positive |- _ => specialize (H r) + end. + congruence. + auto. +Qed. + +Lemma step_instr_block_same : + forall ge sp st st', + step_instr_block ge sp st nil st' -> + st = st'. +Proof. inversion 1; auto. Qed. + +Lemma step_instr_seq_same : + forall ge sp st st', + step_instr_seq ge sp st nil st' -> + st = st'. +Proof. inversion 1; auto. Qed. + +Lemma sem_update' : + forall A ge sp st a x st', + sem ge sp st (update (abstract_sequence empty a) x) st' -> + exists st'', + @step_instr A ge sp st'' x st' /\ + sem ge sp st (abstract_sequence empty a) st''. +Proof. + Admitted. + +Lemma rtlpar_trans_correct : + forall bb ge sp sem_st' sem_st st, + sem ge sp sem_st (abstract_sequence empty (concat (concat bb))) sem_st' -> + match_states sem_st st -> + exists st', RTLPar.step_instr_block ge sp st bb st' + /\ match_states sem_st' st'. +Proof. + induction bb using rev_ind. + { repeat econstructor. eapply abstract_interp_empty3 in H. + inv H. inv H0. constructor; congruence. } + { simplify. inv H0. repeat rewrite concat_app in H. simplify. + rewrite app_nil_r in H. + exploit sem_separate; eauto; inv_simp. + repeat econstructor. admit. admit. + } +Admitted. + +(*Lemma abstract_execution_correct_ld: + forall bb bb' cfi ge tge sp st st' tst, + RTLBlock.step_instr_list ge sp st bb st' -> + ge_preserved ge tge -> + schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi) = true -> + match_states_ld st tst -> + exists tst', RTLPar.step_instr_block tge sp tst bb' tst' + /\ match_states st' tst'. +Proof. + intros.*) +*) + +Lemma match_states_list : + forall A (rs: Regmap.t A) rs', + (forall r, rs !! r = rs' !! r) -> + forall l, rs ## l = rs' ## l. +Proof. induction l; crush. Qed. + +Lemma PTree_matches : + forall A (v: A) res rs rs', + (forall r, rs !! r = rs' !! r) -> + forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x. +Proof. + intros; destruct (Pos.eq_dec x res); subst; + [ repeat rewrite Regmap.gss by auto + | repeat rewrite Regmap.gso by auto ]; auto. +Qed. + +Lemma abstract_interp_empty3 : + forall A ctx st', + @sem A ctx empty st' -> match_states (ctx_is ctx) st'. +Proof. + inversion 1; subst; simplify. destruct ctx. + destruct ctx_is. + constructor; intros. + - inv H0. specialize (H3 x). inv H3. inv H8. reflexivity. + - inv H1. specialize (H3 x). inv H3. inv H8. reflexivity. + - inv H2. inv H8. reflexivity. +Qed. + +Lemma step_instr_matches : + forall A a ge sp st st', + @step_instr A ge sp st a st' -> + forall tst, + match_states st tst -> + exists tst', step_instr ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction 1; simplify; + match goal with H: match_states _ _ |- _ => inv H end; + try solve [repeat econstructor; try erewrite match_states_list; + try apply PTree_matches; eauto; + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end]. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + apply PTree_matches; assumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + econstructor; auto. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + apply PTree_matches; assumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + econstructor; auto. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + erewrite <- eval_predf_pr_equiv; eassumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + erewrite <- eval_predf_pr_equiv; eassumption. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + - admit. Admitted. + +Lemma step_instr_list_matches : + forall a ge sp st st', + step_instr_list ge sp st a st' -> + forall tst, match_states st tst -> + exists tst', step_instr_list ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction a; intros; inv H; + try (exploit step_instr_matches; eauto; []; simplify; + exploit IHa; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma step_instr_seq_matches : + forall a ge sp st st', + step_instr_seq ge sp st a st' -> + forall tst, match_states st tst -> + exists tst', step_instr_seq ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction a; intros; inv H; + try (exploit step_instr_list_matches; eauto; []; simplify; + exploit IHa; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma step_instr_block_matches : + forall bb ge sp st st', + step_instr_block ge sp st bb st' -> + forall tst, match_states st tst -> + exists tst', step_instr_block ge sp tst bb tst' + /\ match_states st' tst'. +Proof. + induction bb; intros; inv H; + try (exploit step_instr_seq_matches; eauto; []; simplify; + exploit IHbb; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma rtlblock_trans_correct' : + forall bb ge sp st x st'', + RTLBlock.step_instr_list ge sp st (bb ++ x :: nil) st'' -> + exists st', RTLBlock.step_instr_list ge sp st bb st' + /\ step_instr ge sp st' x st''. +Proof. + induction bb. + crush. exists st. + split. constructor. inv H. inv H6. auto. + crush. inv H. exploit IHbb. eassumption. simplify. + econstructor. split. + econstructor; eauto. eauto. +Qed. + +Lemma abstract_interp_empty A st : @sem A st empty (ctx_is st). +Proof. destruct st, ctx_is. simpl. repeat econstructor. Qed. + +Lemma abstract_seq : + forall l f i, + abstract_sequence f (l ++ i :: nil) = update (abstract_sequence f l) i. +Proof. induction l; crush. Qed. + +Lemma abstract_sequence_update : + forall l r f, + check_dest_l l r = false -> + (abstract_sequence f l) # (Reg r) = f # (Reg r). +Proof. + induction l using rev_ind; crush. + rewrite abstract_seq. rewrite check_dest_update. apply IHl. + apply check_list_l_false in H. tauto. + apply check_list_l_false in H. tauto. +Qed. + +(*Lemma sem_separate : + forall A ctx b a st', + sem ctx (abstract_sequence empty (a ++ b)) st' -> + exists st'', + @sem A ctx (abstract_sequence empty a) st'' + /\ @sem A (mk_ctx st'' (ctx_sp ctx) (ctx_ge ctx)) (abstract_sequence empty b) st'. +Proof. + induction b using rev_ind; simplify. + { econstructor. simplify. rewrite app_nil_r in H. eauto. apply abstract_interp_empty. } + { simplify. rewrite app_assoc in H. rewrite abstract_seq in H. + exploit sem_update'; eauto; simplify. + exploit IHb; eauto; inv_simp. + econstructor; split; eauto. + rewrite abstract_seq. + eapply sem_update2; eauto. + } +Qed.*) + +Lemma sem_update_RBnop : + forall A ctx f st', + @sem A ctx f st' -> sem ctx (update f RBnop) st'. +Proof. auto. Qed. + +Lemma sem_update_Op : + forall A ge sp ist f st' r l o0 o m rs v ps, + @sem A (mk_ctx ist sp ge) f st' -> + eval_predf ps o = true -> + Op.eval_operation ge sp o0 (rs ## l) m = Some v -> + match_states st' (mk_instr_state rs ps m) -> + exists tst, + sem (mk_ctx ist sp ge) (update f (RBop (Some o) o0 l r)) tst + /\ match_states (mk_instr_state (Regmap.set r v rs) ps m) tst. +Proof. + intros. inv H1. inv H. inv H1. inv H3. simplify. + econstructor. simplify. + { constructor; try constructor; intros; try solve [rewrite genmap1; now eauto]. + destruct (Pos.eq_dec x r); subst. + { rewrite map2. specialize (H1 r). inv H1. +(*} + } + destruct st. + econstructor. simplify. + { constructor. + { constructor. intros. destruct (Pos.eq_dec x r); subst. + { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. + { inv H9. eapply gen_list_base; eauto. } + { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } + { inv H. rewrite genmap1; crush. eauto. } } + { constructor; eauto. intros. + destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed.*) Admitted. + +Lemma sem_update : + forall A ge sp st x st' st'' st''' f, + sem (mk_ctx st sp ge) f st' -> + 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. + 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', + RTLBlock.step_instr_list ge sp st bb st' -> + forall tst, + match_states st tst -> + exists tst', sem (mk_ctx tst sp ge) (abstract_sequence empty bb) tst' + /\ match_states st' tst'. +Proof. + induction bb using rev_ind; simplify. + { econstructor. simplify. apply abstract_interp_empty. + inv H. auto. } + { apply rtlblock_trans_correct' in H. simplify. + rewrite abstract_seq. + exploit IHbb; try eassumption; []; simplify. + exploit sem_update. apply H1. symmetry; eassumption. + eauto. simplify. econstructor. split. apply H3. + auto. } +Qed. + +Lemma abstract_execution_correct: + forall bb bb' cfi cfi' ge tge sp st st' tst, + RTLBlock.step_instr_list ge sp st bb st' -> + ge_preserved ge tge -> + schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi') = true -> + match_states st tst -> + exists tst', RTLPar.step_instr_block tge sp tst bb' tst' + /\ match_states st' tst'. +Proof. + intros. + unfold schedule_oracle in *. simplify. unfold empty_trees in H4. + exploit rtlblock_trans_correct; try eassumption; []; simplify. +(*) exploit abstract_execution_correct'; + try solve [eassumption | apply state_lessdef_match_sem; eassumption]. + apply match_states_commut. eauto. inv_simp. + exploit rtlpar_trans_correct; try eassumption; []; inv_simp. + exploit step_instr_block_matches; eauto. apply match_states_commut; eauto. inv_simp. + repeat match goal with | H: match_states _ _ |- _ => inv H end. + do 2 econstructor; eauto. + econstructor; congruence. +Qed.*)Admitted. + +Definition match_prog (prog : RTLBlock.program) (tprog : RTLPar.program) := + match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq prog tprog. + +Inductive match_stackframes: RTLBlock.stackframe -> RTLPar.stackframe -> Prop := +| match_stackframe: + forall f tf res sp pc rs rs' ps ps', + transl_function f = OK tf -> + (forall x, rs !! x = rs' !! x) -> + (forall x, ps !! x = ps' !! x) -> + match_stackframes (Stackframe res f sp pc rs ps) + (Stackframe res tf sp pc rs' ps'). + +Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop := +| match_state: + forall sf f sp pc rs rs' m sf' tf ps ps' + (TRANSL: transl_function f = OK tf) + (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_returnstate: + forall stack stack' v m + (STACKS: list_forall2 match_stackframes stack stack'), + match_states (Returnstate stack v m) + (Returnstate stack' v m) +| match_callstate: + forall stack stack' f tf args m + (TRANSL: transl_fundef f = OK tf) + (STACKS: list_forall2 match_stackframes stack stack'), + match_states (Callstate stack f args m) + (Callstate stack' tf args m). + +Section CORRECTNESS. + + Context (prog: RTLBlock.program) (tprog : RTLPar.program). + Context (TRANSL: match_prog prog tprog). + + Let ge : RTLBlock.genv := Globalenvs.Genv.globalenv prog. + Let tge : RTLPar.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. + Hint Resolve symbols_preserved : rtlgp. + + Lemma function_ptr_translated: + forall (b: Values.block) (f: RTLBlock.fundef), + Genv.find_funct_ptr ge b = Some f -> + exists tf, + Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = Errors.OK tf. + Proof using TRANSL. + intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. + intros (cu & tf & P & Q & R); exists tf; auto. + Qed. + + Lemma functions_translated: + forall (v: Values.val) (f: RTLBlock.fundef), + Genv.find_funct ge v = Some f -> + exists tf, + Genv.find_funct tge v = Some tf /\ transl_fundef f = Errors.OK tf. + Proof using TRANSL. + intros. exploit (Genv.find_funct_match TRANSL); eauto. + intros (cu & tf & P & Q & R); exists tf; auto. + Qed. + + Lemma senv_preserved: + Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). + Proof (Genv.senv_transf_partial TRANSL). + Hint Resolve senv_preserved : rtlgp. + + Lemma sig_transl_function: + forall (f: RTLBlock.fundef) (tf: RTLPar.fundef), + transl_fundef f = OK tf -> + funsig tf = funsig f. + Proof using . + unfold transl_fundef, transf_partial_fundef, transl_function; intros; + repeat destruct_match; crush; + match goal with H: OK _ = OK _ |- _ => inv H end; auto. + Qed. + Hint Resolve sig_transl_function : rtlgp. + + Hint Resolve Val.lessdef_same : rtlgp. + Hint Resolve regs_lessdef_regs : rtlgp. + + Lemma find_function_translated: + forall ros rs rs' f, + (forall x, rs !! x = rs' !! x) -> + find_function ge ros rs = Some f -> + exists tf, find_function tge ros rs' = Some tf + /\ transl_fundef f = OK tf. + Proof using TRANSL. + Ltac ffts := match goal with + | [ H: forall _, Val.lessdef _ _, r: Registers.reg |- _ ] => + specialize (H r); inv H + | [ H: Vundef = ?r, H1: Genv.find_funct _ ?r = Some _ |- _ ] => + rewrite <- H in H1 + | [ H: Genv.find_funct _ Vundef = Some _ |- _] => solve [inv H] + | _ => solve [exploit functions_translated; eauto] + end. + destruct ros; simplify; try rewrite <- H; + [| rewrite symbols_preserved; destruct_match; + try (apply function_ptr_translated); crush ]; + intros; + repeat ffts. + Qed. + + Lemma schedule_oracle_nil: + forall bb cfi, + schedule_oracle {| bb_body := nil; bb_exit := cfi |} bb = true -> + bb_body bb = nil /\ bb_exit bb = cfi. + Proof using . + unfold schedule_oracle, check_control_flow_instr. + simplify; repeat destruct_match; crush. + Qed. + + Lemma schedule_oracle_nil2: + forall cfi, + schedule_oracle {| bb_body := nil; bb_exit := cfi |} + {| bb_body := nil; bb_exit := cfi |} = true. + Proof using . + unfold schedule_oracle, check_control_flow_instr. + simplify; repeat destruct_match; crush. + Qed. + + Lemma eval_op_eq: + forall (sp0 : Values.val) (op : Op.operation) (vl : list Values.val) m, + Op.eval_operation ge sp0 op vl m = Op.eval_operation tge sp0 op vl m. + Proof using TRANSL. + intros. + destruct op; auto; unfold Op.eval_operation, Genv.symbol_address, Op.eval_addressing32; + [| destruct a; unfold Genv.symbol_address ]; + try rewrite symbols_preserved; auto. + Qed. + Hint Resolve eval_op_eq : rtlgp. + + Lemma eval_addressing_eq: + forall sp addr vl, + Op.eval_addressing ge sp addr vl = Op.eval_addressing tge sp addr vl. + Proof using TRANSL. + intros. + destruct addr; + unfold Op.eval_addressing, Op.eval_addressing32; + unfold Genv.symbol_address; + try rewrite symbols_preserved; auto. + Qed. + Hint Resolve eval_addressing_eq : rtlgp. + + Lemma ge_preserved_lem: + ge_preserved ge tge. + Proof using TRANSL. + unfold ge_preserved. + eauto with rtlgp. + Qed. + Hint Resolve ge_preserved_lem : rtlgp. + + Lemma lessdef_regmap_optget: + forall or rs rs', + regs_lessdef rs rs' -> + Val.lessdef (regmap_optget or Vundef rs) (regmap_optget or Vundef rs'). + Proof using. destruct or; crush. Qed. + Hint Resolve lessdef_regmap_optget : rtlgp. + + Lemma regmap_equiv_lessdef: + forall rs rs', + (forall x, rs !! x = rs' !! x) -> + regs_lessdef rs rs'. + Proof using. + intros; unfold regs_lessdef; intros. + rewrite H. apply Val.lessdef_refl. + Qed. + Hint Resolve regmap_equiv_lessdef : rtlgp. + + Lemma int_lessdef: + forall rs rs', + regs_lessdef rs rs' -> + (forall arg v, + rs !! arg = Vint v -> + rs' !! arg = Vint v). + Proof using. intros ? ? H; intros; specialize (H arg); inv H; crush. Qed. + Hint Resolve int_lessdef : rtlgp. + + Ltac semantics_simpl := + match goal with + | [ H: match_states _ _ |- _ ] => + let H2 := fresh "H" in + learn H as H2; inv H2 + | [ H: transl_function ?f = OK _ |- _ ] => + let H2 := fresh "TRANSL" in + learn H as H2; + unfold transl_function in H2; + destruct (check_scheduled_trees + (@fn_code RTLBlock.bb f) + (@fn_code RTLPar.bb (schedule f))) eqn:?; + [| discriminate ]; inv H2 + | [ H: context[check_scheduled_trees] |- _ ] => + let H2 := fresh "CHECK" in + learn H as H2; + eapply check_scheduled_trees_correct in H2; [| solve [eauto] ] + | [ H: schedule_oracle {| bb_body := nil; bb_exit := _ |} ?bb = true |- _ ] => + let H2 := fresh "SCHED" in + learn H as H2; + apply schedule_oracle_nil in H2 + | [ H: find_function _ _ _ = Some _, H2: forall x, ?rs !! x = ?rs' !! x |- _ ] => + learn H; exploit find_function_translated; try apply H2; eauto; inversion 1 + | [ H: Mem.free ?m _ _ _ = Some ?m', H2: Mem.extends ?m ?m'' |- _ ] => + learn H; exploit Mem.free_parallel_extends; eauto; intros + | [ H: Events.eval_builtin_args _ _ _ _ _ _, H2: regs_lessdef ?rs ?rs' |- _ ] => + let H3 := fresh "H" in + learn H; exploit Events.eval_builtin_args_lessdef; [apply H2 | | |]; + eauto with rtlgp; intro H3; learn H3 + | [ H: Events.external_call _ _ _ _ _ _ _ |- _ ] => + let H2 := fresh "H" in + learn H; exploit Events.external_call_mem_extends; + eauto; intro H2; learn H2 + | [ H: exists _, _ |- _ ] => inv H + | _ => progress simplify + end. + + Hint Resolve Events.eval_builtin_args_preserved : rtlgp. + Hint Resolve Events.external_call_symbols_preserved : rtlgp. + Hint Resolve set_res_lessdef : rtlgp. + Hint Resolve set_reg_lessdef : rtlgp. + Hint Resolve Op.eval_condition_lessdef : rtlgp. + + Hint Constructors Events.eval_builtin_arg: barg. + + Lemma eval_builtin_arg_eq: + forall A ge a v1 m1 e1 e2 sp, + (forall x, e1 x = e2 x) -> + @Events.eval_builtin_arg A ge e1 sp m1 a v1 -> + Events.eval_builtin_arg ge e2 sp m1 a v1. +Proof. induction 2; try rewrite H; eauto with barg. Qed. + + Lemma eval_builtin_args_eq: + forall A ge e1 sp m1 e2 al vl1, + (forall x, e1 x = e2 x) -> + @Events.eval_builtin_args A ge e1 sp m1 al vl1 -> + Events.eval_builtin_args ge e2 sp m1 al vl1. + Proof. + induction 2. + - econstructor; split. + - exploit eval_builtin_arg_eq; eauto. intros. + destruct IHlist_forall2 as [| y]. constructor; eauto. + constructor. constructor; auto. + constructor; eauto. + Qed. + + Lemma step_cf_instr_correct: + forall cfi t s s', + step_cf_instr ge s cfi t s' -> + forall r, + match_states s r -> + exists r', step_cf_instr tge r cfi t r' /\ match_states s' r'. + Proof using TRANSL. + induction 1; repeat semantics_simpl; + try solve [repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp)]. + { do 3 (try erewrite match_states_list by eauto; econstructor; eauto with rtlgp). + eapply eval_builtin_args_eq. eapply REG. + eapply Events.eval_builtin_args_preserved. eapply symbols_preserved. + eauto. + intros. + unfold regmap_setres. destruct res. + destruct (Pos.eq_dec x0 x); subst. + repeat rewrite Regmap.gss; auto. + repeat rewrite Regmap.gso; auto. + eapply REG. eapply REG. + } + { repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). + unfold regmap_optget. destruct or. rewrite REG. constructor; eauto. + constructor; eauto. + } + { exploit IHstep_cf_instr; eauto. simplify. + repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). + erewrite eval_predf_pr_equiv; eauto. + } + Qed. + + Theorem transl_step_correct : + forall (S1 : RTLBlock.state) t S2, + RTLBlock.step ge S1 t S2 -> + forall (R1 : RTLPar.state), + match_states S1 R1 -> + exists R2, Smallstep.plus RTLPar.step tge R1 t R2 /\ match_states S2 R2. + Proof. + + induction 1; repeat semantics_simpl. + + { destruct bb; destruct x. + assert (bb_exit = bb_exit0). + { unfold schedule_oracle in *. simplify. + unfold check_control_flow_instr in *. + destruct_match; crush. + } + subst. + + exploit abstract_execution_correct; try eassumption. eapply ge_preserved_lem. + econstructor; eauto. + simplify. destruct x. inv H7. + + exploit step_cf_instr_correct; try eassumption. econstructor; eauto. + simplify. + + econstructor. econstructor. eapply Smallstep.plus_one. econstructor. + 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. + 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. } + { inv STACKS. inv H2. repeat econstructor; eauto. + intros. apply PTree_matches; eauto. } + Qed. + + Lemma transl_initial_states: + forall S, + RTLBlock.initial_state prog S -> + exists R, RTLPar.initial_state tprog R /\ match_states S R. + Proof. + induction 1. + exploit function_ptr_translated; eauto. intros [tf [A B]]. + econstructor; split. + econstructor. apply (Genv.init_mem_transf_partial TRANSL); eauto. + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto. + symmetry; eapply match_program_main; eauto. + eexact A. + rewrite <- H2. apply sig_transl_function; auto. + constructor. auto. constructor. + Qed. + + Lemma transl_final_states: + forall S R r, + match_states S R -> RTLBlock.final_state S r -> RTLPar.final_state R r. + Proof. + intros. inv H0. inv H. inv STACKS. constructor. + Qed. + + Theorem transf_program_correct: + Smallstep.forward_simulation (RTLBlock.semantics prog) (RTLPar.semantics tprog). + Proof. + eapply Smallstep.forward_simulation_plus. + apply senv_preserved. + eexact transl_initial_states. + eexact transl_final_states. + exact transl_step_correct. + Qed. + +End CORRECTNESS. +#+end_src + +* License + +#+name: license +#+begin_src coq :tangle no +(* + * 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 . + *) +#+end_src 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 --- Makefile | 5 +- lit/scheduling.org | 302 ++++++++++++++++++++++++++---------------------- src/hls/RTLBlockInstr.v | 49 +++----- src/hls/RTLBlockgen.v | 26 +++-- 4 files changed, 199 insertions(+), 183 deletions(-) diff --git a/Makefile b/Makefile index 3ff20a6..3b51d00 100644 --- a/Makefile +++ b/Makefile @@ -87,7 +87,10 @@ clean:: Makefile.coq rm -f Makefile.coq detangle-all: - emacs --batch --eval "(progn(require 'org)(require 'ob-tangle)$(foreach vs,$(VS),(org-babel-detangle \"$(vs)\")))" + emacs --batch --eval "(progn(require 'org)(require 'ob-tangle)\ + (setq org-src-preserve-indentation t)\ + $(foreach vs,$(VS),(org-babel-detangle \"$(vs)\"))\ + (org-save-all-org-buffers))" clean:: rm -f */*.v.d */*.glob */*.vo */*~ *~ diff --git a/lit/scheduling.org b/lit/scheduling.org index 7970019..c3fb12a 100644 --- a/lit/scheduling.org +++ b/lit/scheduling.org @@ -13,7 +13,7 @@ <> #+end_src -#+name: rtlblockinstr-main +#+name: rtlblockinstr-imports #+begin_src coq Require Import Coq.micromega.Lia. @@ -29,20 +29,20 @@ Require Import compcert.verilog.Op. Require Import Predicate. Require Import Vericertlib. +#+end_src -(** * 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. +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]). -*) +#+name: rtlblockinstr-instr-def +#+begin_src coq Definition node := positive. Inductive instr : Type := @@ -51,12 +51,15 @@ 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. +#+end_src -(** ** 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. +#+name: rtlblockinstr-cf-instr-def +#+begin_src coq Inductive cf_instr : Type := | RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr | RBtailcall : signature -> reg + ident -> list reg -> cf_instr @@ -67,9 +70,12 @@ Inductive cf_instr : Type := | RBreturn : option reg -> cf_instr | RBgoto : node -> cf_instr | RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. +#+end_src -(** ** Helper functions *) +*** Helper functions +#+name: rtlblockinstr-helpers +#+begin_src coq Fixpoint successors_instr (i : cf_instr) : list node := match i with | RBcall sig ros args res s => s :: nil @@ -158,24 +164,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. +#+end_src -(** *** 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. +#+name: rtlblockinstr-instr-state +#+begin_src coq Record instr_state := mk_instr_state { is_rs: regset; is_ps: predset; is_mem: mem; }. +#+end_src -(** ** Top-Level Type Definitions *) +*** Top-Level Type Definitions +#+name: rtlblockinstr-type-def +#+begin_src coq Section DEFINITION. Context {bblock_body: Type}. @@ -238,9 +250,12 @@ Section DEFINITION. state. End DEFINITION. +#+end_src -(** ** Semantics *) +*** Semantics +#+name: rtlblockinstr-semantics +#+begin_src coq Section RELSEM. Context {bblock_body : Type}. @@ -1540,7 +1555,7 @@ Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]]. <> #+end_src -#+name: rtlblockgen-main +#+name: rtlblockgen-imports #+begin_src coq Require compcert.backend.RTL. Require Import compcert.common.AST. @@ -1553,7 +1568,15 @@ Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLBlock. #[local] Open Scope positive. +#+end_src +#+name: rtlblockgen-equalities-insert +#+begin_src coq +<> +#+end_src + +#+name: rtlblockgen-main +#+begin_src coq 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 @@ -1563,124 +1586,6 @@ Definition find_block (max: positive) (nodes: list positive) (index: positive) : (*Compute find_block (2::94::28::40::19::nil) 40.*) -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. - Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) := match istr, istr' with | RTL.Inop n', RBnop => (n' + 1 =? n) @@ -1812,6 +1717,129 @@ Definition transl_program : RTL.program -> Errors.res program := transform_partial_program transl_fundef. #+end_src +** Equalities (rtlblockgen-equalities) + +#+name: rtlblockgen-equalities +#+begin_src coq :tangle no +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. +#+end_src + * RTLBlockgenproof :PROPERTIES: :header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v 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 --- lit/scheduling.org | 241 ++++++++++++++++++++++++-------------------------- src/hls/RTLBlockgen.v | 2 - 2 files changed, 117 insertions(+), 126 deletions(-) diff --git a/lit/scheduling.org b/lit/scheduling.org index c3fb12a..728dfc5 100644 --- a/lit/scheduling.org +++ b/lit/scheduling.org @@ -1572,7 +1572,123 @@ Require Import vericert.hls.RTLBlock. #+name: rtlblockgen-equalities-insert #+begin_src coq -<> +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. #+end_src #+name: rtlblockgen-main @@ -1717,129 +1833,6 @@ Definition transl_program : RTL.program -> Errors.res program := transform_partial_program transl_fundef. #+end_src -** Equalities (rtlblockgen-equalities) - -#+name: rtlblockgen-equalities -#+begin_src coq :tangle no -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. -#+end_src - * RTLBlockgenproof :PROPERTIES: :header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v 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 6b2e01e97ca67219465b92e0fbe536bf2a5434e7 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 23 Mar 2022 09:21:41 +0000 Subject: Create a few more org files about specific topics --- lit/basic-block-generation.org | 378 +++++ lit/scheduler-languages.org | 689 +++++++++ lit/scheduler.org | 2295 ++++++++++++++++++++++++++++ lit/scheduling.org | 3303 ---------------------------------------- 4 files changed, 3362 insertions(+), 3303 deletions(-) create mode 100644 lit/basic-block-generation.org create mode 100644 lit/scheduler-languages.org create mode 100644 lit/scheduler.org delete mode 100644 lit/scheduling.org diff --git a/lit/basic-block-generation.org b/lit/basic-block-generation.org new file mode 100644 index 0000000..4a3edc4 --- /dev/null +++ b/lit/basic-block-generation.org @@ -0,0 +1,378 @@ +#+title: Basic Block Generation +#+author: Yann Herklotz +#+email: yann [at] yannherklotz [dot] com + +* RTLBlockgen +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgen.v +:END: + +Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]]. + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblockgen-imports +#+begin_src coq +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. +#+end_src + +#+name: rtlblockgen-equalities-insert +#+begin_src coq +<> +#+end_src + +#+name: rtlblockgen-main +#+begin_src coq +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 + | 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. +#+end_src + +** Equalities + +#+name: rtlblockgen-equalities +#+begin_src coq :tangle no +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. +#+end_src + +* RTLBlockgenproof +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +** Imports + +#+name: rtlblockgenproof-imports +#+begin_src coq +Require compcert.backend.RTL. +Require Import compcert.common.AST. +Require Import compcert.lib.Maps. + +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLBlockgen. +#+end_src + +** Match states + +The ~match_states~ predicate describes which states are equivalent between the two languages, in this +case ~RTL~ and ~RTLBlock~. + +#+name: rtlblockgenproof-match-states +#+begin_src coq +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). +#+end_src + +** Correctness + +#+name: rtlblockgenproof-correctness +#+begin_src coq +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. +#+end_src + +* License + +#+name: license +#+begin_src coq :tangle no +(* + * 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 . + *) +#+end_src diff --git a/lit/scheduler-languages.org b/lit/scheduler-languages.org new file mode 100644 index 0000000..cf03b3a --- /dev/null +++ b/lit/scheduler-languages.org @@ -0,0 +1,689 @@ +#+title: Scheduler Languages +#+author: Yann Herklotz +#+email: yann [at] yannherklotz [dot] com + +* RTLBlockInstr +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockInstr.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblockinstr-imports +#+begin_src coq +Require Import Coq.micromega.Lia. + +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.Values. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. +Require Import compcert.verilog.Op. + +Require Import Predicate. +Require Import Vericertlib. +#+end_src + +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]). + +#+name: rtlblockinstr-instr-def +#+begin_src coq +Definition node := positive. + +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. +#+end_src + +** Control-Flow Instruction Definition + +These are the instructions that count as control-flow, and will be placed at the end of the basic +blocks. + +#+name: rtlblockinstr-cf-instr-def +#+begin_src coq +Inductive cf_instr : Type := +| RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr +| RBtailcall : signature -> reg + ident -> list reg -> cf_instr +| RBbuiltin : external_function -> list (builtin_arg reg) -> + builtin_res reg -> node -> cf_instr +| RBcond : condition -> list reg -> node -> node -> cf_instr +| RBjumptable : reg -> list node -> cf_instr +| RBreturn : option reg -> cf_instr +| RBgoto : node -> cf_instr +| RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. +#+end_src + +** Helper functions + +#+name: rtlblockinstr-helpers +#+begin_src coq +Fixpoint successors_instr (i : cf_instr) : list node := + match i with + | RBcall sig ros args res s => s :: nil + | RBtailcall sig ros args => nil + | RBbuiltin ef args res s => s :: nil + | RBcond cond args ifso ifnot => ifso :: ifnot :: nil + | RBjumptable arg tbl => tbl + | RBreturn optarg => nil + | RBgoto n => n :: nil + | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil) + end. + +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) + | RBload p chunk addr args dst => + fold_left Pos.max args (Pos.max dst m) + | RBstore p chunk addr args src => + fold_left Pos.max args (Pos.max src m) + | RBsetpred p' c args p => + 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)) + | RBcall sig (inr id) args res s => + fold_left Pos.max args (Pos.max res m) + | RBtailcall sig (inl r) args => + fold_left Pos.max args (Pos.max r m) + | RBtailcall sig (inr id) args => + 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) + | RBcond cond args ifso ifnot => fold_left Pos.max args m + | RBjumptable arg tbl => Pos.max arg m + | RBreturn None => m + | RBreturn (Some arg) => Pos.max arg m + | RBgoto n => m + | RBpred_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. +#+end_src + +*** 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. + +#+name: rtlblockinstr-instr-state +#+begin_src coq +Record instr_state := mk_instr_state { + is_rs: regset; + is_ps: predset; + is_mem: mem; +}. +#+end_src + +** Top-Level Type Definitions + +#+name: rtlblockinstr-type-def +#+begin_src coq +Section DEFINITION. + + Context {bblock_body: Type}. + + 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_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. + +End DEFINITION. +#+end_src + +** Semantics + +#+name: rtlblockinstr-semantics +#+begin_src coq +Section RELSEM. + + Context {bblock_body : Type}. + + Definition genv := Genv.t (@fundef bblock_body) unit. + + 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_RBnop: + 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 + | 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 + | 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 + | 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. + + 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, + 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') + 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) (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') + | 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) + | 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) + | 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) + 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) + | 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'. + +End RELSEM. +#+end_src + +* RTLBlock +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlock.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblock-main +#+begin_src coq +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.RTLBlockInstr. + +Definition bb := list instr. + +Definition bblock := @bblock bb. +Definition code := @code bb. +Definition function := @function bb. +Definition fundef := @fundef bb. +Definition program := @program bb. +Definition funsig := @funsig bb. +Definition stackframe := @stackframe 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). + + 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_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). +#+end_src + +* RTLPar +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPar.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpar-main +#+begin_src coq +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.RTLBlockInstr. + +Definition bb := list (list (list instr)). + +Definition bblock := @bblock bb. +Definition code := @code bb. +Definition function := @function bb. +Definition fundef := @fundef bb. +Definition program := @program bb. +Definition funsig := @funsig bb. +Definition stackframe := @stackframe bb. +Definition state := @state bb. +Definition genv := @genv bb. + +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_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 -> bb -> 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: 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 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') + | 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. +#+end_src + +* License + +#+name: license +#+begin_src coq :tangle no +(* + * 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 . + *) +#+end_src diff --git a/lit/scheduler.org b/lit/scheduler.org new file mode 100644 index 0000000..018633c --- /dev/null +++ b/lit/scheduler.org @@ -0,0 +1,2295 @@ +#+title: Basic Block Generation +#+author: Yann Herklotz +#+email: yann [at] yannherklotz [dot] com + +* Scheduler +:PROPERTIES: +:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Schedule.ml +:END: + +#+begin_src ocaml :comments no :padline no :exports none +<> +#+end_src + +#+name: scheduler-main +#+begin_src ocaml +open Printf +open Clflags +open Camlcoq +open Datatypes +open Coqlib +open Maps +open AST +open Kildall +open Op +open RTLBlockInstr +open Predicate +open RTLBlock +open HTL +open Verilog +open HTLgen +open HTLMonad +open HTLMonadExtra + +module SS = Set.Make(P) + +type svtype = + | BBType of int + | VarType of int * int + +type sv = { + sv_type: svtype; + sv_num: int; +} + +let print_sv v = + match v with + | { sv_type = BBType bbi; sv_num = n } -> sprintf "bb%d_%d" bbi n + | { sv_type = VarType (bbi, i); sv_num = n } -> sprintf "var%dn%d_%d" bbi i n + +module G = Graph.Persistent.Digraph.ConcreteLabeled(struct + type t = sv + let compare = compare + let equal = (=) + let hash = Hashtbl.hash +end)(struct + type t = int + let compare = compare + let hash = Hashtbl.hash + let equal = (=) + let default = 0 +end) + +module GDot = Graph.Graphviz.Dot(struct + let graph_attributes _ = [] + let default_vertex_attributes _ = [] + let vertex_name = print_sv + let vertex_attributes _ = [] + let get_subgraph _ = None + let default_edge_attributes _ = [] + let edge_attributes _ = [] + + include G + end) + +module DFG = Graph.Persistent.Digraph.ConcreteLabeled(struct + type t = int * instr + let compare = compare + let equal = (=) + let hash = Hashtbl.hash +end)(struct + type t = int + let compare = compare + let hash = Hashtbl.hash + let equal = (=) + let default = 0 +end) + +module DFGSimp = Graph.Persistent.Graph.Concrete(struct + type t = int * instr + let compare = compare + let equal = (=) + let hash = Hashtbl.hash + end) + +let convert dfg = + DFG.fold_vertex (fun v g -> DFGSimp.add_vertex g v) dfg DFGSimp.empty + |> DFG.fold_edges (fun v1 v2 g -> DFGSimp.add_edge (DFGSimp.add_edge g v1 v2) v2 v1) dfg + +let reg r = sprintf "r%d" (P.to_int r) +let print_pred r = sprintf "p%d" (P.to_int r) + +let print_instr = function + | RBnop -> "" + | RBload (_, _, _, _, r) -> sprintf "load(%s)" (reg r) + | RBstore (_, _, _, _, r) -> sprintf "store(%s)" (reg r) + | RBsetpred (_, _, _, p) -> sprintf "setpred(%s)" (print_pred p) + | RBop (_, op, args, d) -> + (match op, args with + | Omove, _ -> "mov" + | Ointconst n, _ -> sprintf "%s=%ld" (reg d) (camlint_of_coqint n) + | Olongconst n, _ -> sprintf "%s=%LdL" (reg d) (camlint64_of_coqint n) + | Ofloatconst n, _ -> sprintf "%s=%.15F" (reg d) (camlfloat_of_coqfloat n) + | Osingleconst n, _ -> sprintf "%s=%.15Ff" (reg d) (camlfloat_of_coqfloat32 n) + | Oindirectsymbol id, _ -> sprintf "%s=&%s" (reg d) (extern_atom id) + | Ocast8signed, [r1] -> sprintf "%s=int8signed(%s)" (reg d) (reg r1) + | Ocast8unsigned, [r1] -> sprintf "%s=int8unsigned(%s)" (reg d) (reg r1) + | Ocast16signed, [r1] -> sprintf "%s=int16signed(%s)" (reg d) (reg r1) + | Ocast16unsigned, [r1] -> sprintf "%s=int16unsigned(%s)" (reg d) (reg r1) + | Oneg, [r1] -> sprintf "%s=-%s" (reg d) (reg r1) + | Osub, [r1;r2] -> sprintf "%s=%s-%s" (reg d) (reg r1) (reg r2) + | Omul, [r1;r2] -> sprintf "%s=%s*%s" (reg d) (reg r1) (reg r2) + | Omulimm n, [r1] -> sprintf "%s=%s*%ld" (reg d) (reg r1) (camlint_of_coqint n) + | Omulhs, [r1;r2] -> sprintf "%s=mulhs(%s,%s)" (reg d) (reg r1) (reg r2) + | Omulhu, [r1;r2] -> sprintf "%s=mulhu(%s,%s)" (reg d) (reg r1) (reg r2) + | Odiv, [r1;r2] -> sprintf "%s=%s /s %s" (reg d) (reg r1) (reg r2) + | Odivu, [r1;r2] -> sprintf "%s=%s /u %s" (reg d) (reg r1) (reg r2) + | Omod, [r1;r2] -> sprintf "%s=%s %%s %s" (reg d) (reg r1) (reg r2) + | Omodu, [r1;r2] -> sprintf "%s=%s %%u %s" (reg d) (reg r1) (reg r2) + | Oand, [r1;r2] -> sprintf "%s=%s & %s" (reg d) (reg r1) (reg r2) + | Oandimm n, [r1] -> sprintf "%s=%s & %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oor, [r1;r2] -> sprintf "%s=%s | %s" (reg d) (reg r1) (reg r2) + | Oorimm n, [r1] -> sprintf "%s=%s | %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oxor, [r1;r2] -> sprintf "%s=%s ^ %s" (reg d) (reg r1) (reg r2) + | Oxorimm n, [r1] -> sprintf "%s=%s ^ %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Onot, [r1] -> sprintf "%s=not(%s)" (reg d) (reg r1) + | Oshl, [r1;r2] -> sprintf "%s=%s << %s" (reg d) (reg r1) (reg r2) + | Oshlimm n, [r1] -> sprintf "%s=%s << %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshr, [r1;r2] -> sprintf "%s=%s >>s %s" (reg d) (reg r1) (reg r2) + | Oshrimm n, [r1] -> sprintf "%s=%s >>s %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrximm n, [r1] -> sprintf "%s=%s >>x %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshru, [r1;r2] -> sprintf "%s=%s >>u %s" (reg d) (reg r1) (reg r2) + | Oshruimm n, [r1] -> sprintf "%s=%s >>u %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Ororimm n, [r1] -> sprintf "%s=%s ror %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshldimm n, [r1;r2] -> sprintf "%s=(%s, %s) << %ld" (reg d) (reg r1) (reg r2) (camlint_of_coqint n) + | Olea addr, args -> sprintf "%s=addr" (reg d) + | Omakelong, [r1;r2] -> sprintf "%s=makelong(%s,%s)" (reg d) (reg r1) (reg r2) + | Olowlong, [r1] -> sprintf "%s=lowlong(%s)" (reg d) (reg r1) + | Ohighlong, [r1] -> sprintf "%s=highlong(%s)" (reg d) (reg r1) + | Ocast32signed, [r1] -> sprintf "%s=long32signed(%s)" (reg d) (reg r1) + | Ocast32unsigned, [r1] -> sprintf "%s=long32unsigned(%s)" (reg d) (reg r1) + | Onegl, [r1] -> sprintf "%s=-l %s" (reg d) (reg r1) + | Osubl, [r1;r2] -> sprintf "%s=%s -l %s" (reg d) (reg r1) (reg r2) + | Omull, [r1;r2] -> sprintf "%s=%s *l %s" (reg d) (reg r1) (reg r2) + | Omullimm n, [r1] -> sprintf "%s=%s *l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Omullhs, [r1;r2] -> sprintf "%s=mullhs(%s,%s)" (reg d) (reg r1) (reg r2) + | Omullhu, [r1;r2] -> sprintf "%s=mullhu(%s,%s)" (reg d) (reg r1) (reg r2) + | Odivl, [r1;r2] -> sprintf "%s=%s /ls %s" (reg d) (reg r1) (reg r2) + | Odivlu, [r1;r2] -> sprintf "%s=%s /lu %s" (reg d) (reg r1) (reg r2) + | Omodl, [r1;r2] -> sprintf "%s=%s %%ls %s" (reg d) (reg r1) (reg r2) + | Omodlu, [r1;r2] -> sprintf "%s=%s %%lu %s" (reg d) (reg r1) (reg r2) + | Oandl, [r1;r2] -> sprintf "%s=%s &l %s" (reg d) (reg r1) (reg r2) + | Oandlimm n, [r1] -> sprintf "%s=%s &l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Oorl, [r1;r2] -> sprintf "%s=%s |l %s" (reg d) (reg r1) (reg r2) + | Oorlimm n, [r1] -> sprintf "%s=%s |l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Oxorl, [r1;r2] -> sprintf "%s=%s ^l %s" (reg d) (reg r1) (reg r2) + | Oxorlimm n, [r1] -> sprintf "%s=%s ^l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Onotl, [r1] -> sprintf "%s=notl(%s)" (reg d) (reg r1) + | Oshll, [r1;r2] -> sprintf "%s=%s < sprintf "%s=%s < sprintf "%s=%s >>ls %s" (reg d) (reg r1) (reg r2) + | Oshrlimm n, [r1] -> sprintf "%s=%s >>ls %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrxlimm n, [r1] -> sprintf "%s=%s >>lx %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrlu, [r1;r2] -> sprintf "%s=%s >>lu %s" (reg d) (reg r1) (reg r2) + | Oshrluimm n, [r1] -> sprintf "%s=%s >>lu %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Ororlimm n, [r1] -> sprintf "%s=%s rorl %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oleal addr, args -> sprintf "%s=addr" (reg d) + | Onegf, [r1] -> sprintf "%s=negf(%s)" (reg d) (reg r1) + | Oabsf, [r1] -> sprintf "%s=absf(%s)" (reg d) (reg r1) + | Oaddf, [r1;r2] -> sprintf "%s=%s +f %s" (reg d) (reg r1) (reg r2) + | Osubf, [r1;r2] -> sprintf "%s=%s -f %s" (reg d) (reg r1) (reg r2) + | Omulf, [r1;r2] -> sprintf "%s=%s *f %s" (reg d) (reg r1) (reg r2) + | Odivf, [r1;r2] -> sprintf "%s=%s /f %s" (reg d) (reg r1) (reg r2) + | Onegfs, [r1] -> sprintf "%s=negfs(%s)" (reg d) (reg r1) + | Oabsfs, [r1] -> sprintf "%s=absfs(%s)" (reg d) (reg r1) + | Oaddfs, [r1;r2] -> sprintf "%s=%s +fs %s" (reg d) (reg r1) (reg r2) + | Osubfs, [r1;r2] -> sprintf "%s=%s -fs %s" (reg d) (reg r1) (reg r2) + | Omulfs, [r1;r2] -> sprintf "%s=%s *fs %s" (reg d) (reg r1) (reg r2) + | Odivfs, [r1;r2] -> sprintf "%s=%s /fs %s" (reg d) (reg r1) (reg r2) + | Osingleoffloat, [r1] -> sprintf "%s=singleoffloat(%s)" (reg d) (reg r1) + | Ofloatofsingle, [r1] -> sprintf "%s=floatofsingle(%s)" (reg d) (reg r1) + | Ointoffloat, [r1] -> sprintf "%s=intoffloat(%s)" (reg d) (reg r1) + | Ofloatofint, [r1] -> sprintf "%s=floatofint(%s)" (reg d) (reg r1) + | Ointofsingle, [r1] -> sprintf "%s=intofsingle(%s)" (reg d) (reg r1) + | Osingleofint, [r1] -> sprintf "%s=singleofint(%s)" (reg d) (reg r1) + | Olongoffloat, [r1] -> sprintf "%s=longoffloat(%s)" (reg d) (reg r1) + | Ofloatoflong, [r1] -> sprintf "%s=floatoflong(%s)" (reg d) (reg r1) + | Olongofsingle, [r1] -> sprintf "%s=longofsingle(%s)" (reg d) (reg r1) + | Osingleoflong, [r1] -> sprintf "%s=singleoflong(%s)" (reg d) (reg r1) + | Ocmp c, args -> sprintf "%s=cmp" (reg d) + | Osel (c, ty), r1::r2::args -> sprintf "%s=sel" (reg d) + | _, _ -> sprintf "N/a") + +module DFGDot = Graph.Graphviz.Dot(struct + let graph_attributes _ = [] + let default_vertex_attributes _ = [] + let vertex_name = function (i, instr) -> sprintf "\"%d:%s\"" i (print_instr instr) + let vertex_attributes _ = [] + let get_subgraph _ = None + let default_edge_attributes _ = [] + let edge_attributes _ = [] + + include DFG + end) + +module DFGDfs = Graph.Traverse.Dfs(DFG) + +module IMap = Map.Make (struct + type t = int + + let compare = compare +end) + +let gen_vertex instrs i = (i, List.nth instrs i) + +(** The DFG type defines a list of instructions with their data dependencies as [edges], which are + the pairs of integers that represent the index of the instruction in the [nodes]. The edges + always point from left to right. *) + +let print_list f out_chan a = + fprintf out_chan "[ "; + List.iter (fprintf out_chan "%a " f) a; + fprintf out_chan "]" + +let print_tuple out_chan a = + let l, r = a in + fprintf out_chan "(%d,%d)" l r + +(*let print_dfg out_chan dfg = + fprintf out_chan "{ nodes = %a, edges = %a }" + (print_list PrintRTLBlockInstr.print_bblock_body) + dfg.nodes (print_list print_tuple) dfg.edges*) + +let print_dfg = DFGDot.output_graph + +let read_process command = + let buffer_size = 2048 in + let buffer = Buffer.create buffer_size in + let string = Bytes.create buffer_size in + let in_channel = Unix.open_process_in command in + let chars_read = ref 1 in + while !chars_read <> 0 do + chars_read := input in_channel string 0 buffer_size; + Buffer.add_substring buffer (Bytes.to_string string) 0 !chars_read + done; + ignore (Unix.close_process_in in_channel); + Buffer.contents buffer + +let comb_delay = function + | RBnop -> 0 + | RBop (_, op, _, _) -> + (match op with + | Omove -> 0 + | Ointconst _ -> 0 + | Olongconst _ -> 0 + | Ocast8signed -> 0 + | Ocast8unsigned -> 0 + | Ocast16signed -> 0 + | Ocast16unsigned -> 0 + | Oneg -> 0 + | Onot -> 0 + | Oor -> 0 + | Oorimm _ -> 0 + | Oand -> 0 + | Oandimm _ -> 0 + | Oxor -> 0 + | Oxorimm _ -> 0 + | Omul -> 8 + | Omulimm _ -> 8 + | Omulhs -> 8 + | Omulhu -> 8 + | Odiv -> 72 + | Odivu -> 72 + | Omod -> 72 + | Omodu -> 72 + | _ -> 1) + | _ -> 1 + +let pipeline_stages = function + | RBload _ -> 2 + | RBop (_, op, _, _) -> + (match op with + | Odiv -> 32 + | Odivu -> 32 + | Omod -> 32 + | Omodu -> 32 + | _ -> 0) + | _ -> 0 + +(** Add a dependency if it uses a register that was written to previously. *) +let add_dep map i tree dfg curr = + match PTree.get curr tree with + | None -> dfg + | Some ip -> + let ipv = (List.nth map ip) in + 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. + + This function only gathers the RAW constraints, and will therefore only be active for operations + that modify registers, which is this case only affects loads and operations. *) +let accumulate_RAW_deps map dfg curr = + let i, dst_map, graph = dfg in + let acc_dep_instruction rs dst = + ( i + 1, + PTree.set dst i dst_map, + List.fold_left (add_dep map i dst_map) graph rs + ) + in + let acc_dep_instruction_nodst rs = + ( i + 1, + dst_map, + List.fold_left (add_dep map i dst_map) graph rs) + in + 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) + +(** Finds the next write to the [dst] register. This is a small optimisation so that only one + dependency is generated for a data dependency. *) +let rec find_next_dst_write i dst i' curr = + let check_dst dst' curr' = + if dst = dst' then Some (i, i') + else find_next_dst_write i dst (i' + 1) curr' + in + match curr with + | [] -> None + | RBop (_, _, _, dst') :: curr' -> check_dst dst' curr' + | RBload (_, _, _, _, dst') :: curr' -> check_dst dst' curr' + | _ :: curr' -> find_next_dst_write i dst (i' + 1) curr' + +let rec find_all_next_dst_read i dst i' curr = + let check_dst rs curr' = + if List.exists (fun x -> x = dst) rs + then (i, i') :: find_all_next_dst_read i dst (i' + 1) curr' + else find_all_next_dst_read i dst (i' + 1) curr' + in + match curr with + | [] -> [] + | RBop (_, _, rs, _) :: curr' -> check_dst rs curr' + | RBload (_, _, _, rs, _) :: curr' -> check_dst rs curr' + | RBstore (_, _, _, rs, src) :: curr' -> check_dst (src :: rs) curr' + | RBnop :: curr' -> find_all_next_dst_read i dst (i' + 1) curr' + | RBsetpred (_, _, rs, _) :: curr' -> check_dst rs curr' + +let drop i lst = + let rec drop' i' lst' = + match lst' with + | _ :: ls -> if i' = i then ls else drop' (i' + 1) ls + | [] -> [] + in + if i = 0 then lst else drop' 1 lst + +let take i lst = + let rec take' i' lst' = + match lst' with + | l :: ls -> if i' = i then [ l ] else l :: take' (i' + 1) ls + | [] -> [] + in + if i = 0 then [] else take' 1 lst + +let rec next_store i = function + | [] -> None + | RBstore (_, _, _, _, _) :: _ -> Some i + | _ :: rst -> next_store (i + 1) rst + +let rec next_load i = function + | [] -> None + | RBload (_, _, _, _, _) :: _ -> Some i + | _ :: rst -> next_load (i + 1) rst + +let accumulate_RAW_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBload (_, _, _, _, _) -> ( + 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) ) + | _ -> dfg + +let accumulate_WAR_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBstore (_, _, _, _, _) -> ( + match next_load 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) ) + | _ -> dfg + +let accumulate_WAW_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | 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)) + | _ -> dfg + +(** Predicate dependencies. *) + +let rec in_predicate p p' = + match p' with + | Plit p'' -> P.to_int p = P.to_int (snd p'') + | Pand (p1, p2) -> in_predicate p p1 || in_predicate p p2 + | Por (p1, p2) -> in_predicate p p1 || in_predicate p p2 + | Ptrue -> false + | Pfalse -> false + +let get_predicate = function + | RBop (p, _, _, _) -> p + | RBload (p, _, _, _, _) -> p + | RBstore (p, _, _, _, _) -> p + | RBsetpred (p, _, _, _) -> p + | _ -> None + +let rec next_setpred p i = function + | [] -> None + | RBsetpred (_, _, _, p') :: rst -> + if in_predicate p' p then + Some i + else + next_setpred p (i + 1) rst + | _ :: rst -> next_setpred p (i + 1) rst + +let rec next_preduse p i instr= + let next p' rst = + if in_predicate p p' then + Some i + else + next_preduse p (i + 1) rst + in + match instr with + | [] -> None + | 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 = + let i, curr = curri in + match get_predicate curr with + | Some p -> ( + match next_setpred p 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) ) + | _ -> dfg + +let accumulate_WAR_pred_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBsetpred (_, _, _, p) -> ( + match next_preduse p 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) ) + | _ -> dfg + +let accumulate_WAW_pred_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBsetpred (_, _, _, p) -> ( + match next_setpred (Plit (true, p)) 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) ) + | _ -> dfg + +(** This function calculates the WAW dependencies, which happen when two writes are ordered one + after another and therefore have to be kept in that order. This accumulation might be redundant + if register renaming is done before hand, because then these dependencies can be avoided. *) +let accumulate_WAW_deps instrs dfg curri = + let i, curr = curri in + let dst_dep dst = + match find_next_dst_write i dst (i + 1) (drop (i + 1) instrs) with + | Some (a, b) -> DFG.add_edge dfg (gen_vertex instrs a) (gen_vertex instrs b) + | _ -> dfg + in + match curr with + | RBop (_, _, _, dst) -> dst_dep dst + | RBload (_, _, _, _, dst) -> dst_dep dst + | RBstore (_, _, _, _, _) -> ( + match next_store (i + 1) (drop (i + 1) instrs) with + | None -> dfg + | Some i' -> DFG.add_edge dfg (gen_vertex instrs i) (gen_vertex instrs i') ) + | _ -> dfg + +let accumulate_WAR_deps instrs dfg curri = + let i, curr = curri in + let dst_dep dst = + let dep_list = find_all_next_dst_read i dst 0 (take i instrs |> List.rev) + |> List.map (function (d, d') -> (i - d' - 1, d)) + in + List.fold_left (fun g -> + function (d, d') -> DFG.add_edge g (gen_vertex instrs d) (gen_vertex instrs d')) dfg dep_list + in + match curr with + | RBop (_, _, _, dst) -> dst_dep dst + | RBload (_, _, _, _, dst) -> dst_dep dst + | _ -> dfg + +let assigned_vars vars = function + | RBnop -> vars + | RBop (_, _, _, dst) -> dst :: vars + | RBload (_, _, _, _, dst) -> dst :: vars + | RBstore (_, _, _, _, _) -> vars + | RBsetpred (_, _, _, _) -> vars + +let get_pred = function + | RBnop -> None + | RBop (op, _, _, _) -> op + | RBload (op, _, _, _, _) -> op + | RBstore (op, _, _, _, _) -> op + | RBsetpred (_, _, _, _) -> None + +let independant_pred p p' = + match sat_pred_simple (Pand (p, p')) with + | None -> true + | _ -> false + +let check_dependent op1 op2 = + match op1, op2 with + | Some p, Some p' -> not (independant_pred p p') + | _, _ -> true + +let remove_unnecessary_deps graph = + let is_dependent v1 v2 g = + let (_, instr1) = v1 in + let (_, instr2) = v2 in + if check_dependent (get_pred instr1) (get_pred instr2) + then g + else DFG.remove_edge g v1 v2 + in + DFG.fold_edges is_dependent graph graph + +(** All the nodes in the DFG have to come after the source of the basic block, and should terminate + before the sink of the basic block. After that, there should be constraints for data + dependencies between nodes. *) +let gather_bb_constraints debug bb = + let ibody = List.mapi (fun i a -> (i, a)) bb.bb_body in + let dfg = List.fold_left (fun dfg v -> DFG.add_vertex dfg v) DFG.empty ibody in + let _, _, dfg' = + List.fold_left (accumulate_RAW_deps ibody) + (0, PTree.empty, dfg) + bb.bb_body + in + let dfg'' = List.fold_left (fun dfg f -> List.fold_left (f bb.bb_body) dfg ibody) dfg' + [ accumulate_WAW_deps; + accumulate_WAR_deps; + accumulate_RAW_mem_deps; + accumulate_WAR_mem_deps; + accumulate_WAW_mem_deps; + accumulate_RAW_pred_deps; + accumulate_WAR_pred_deps; + accumulate_WAW_pred_deps + ] + in + let dfg''' = remove_unnecessary_deps dfg'' in + (List.length bb.bb_body, dfg''', successors_instr bb.bb_exit) + +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 + then G.add_edge_e g (encode_bb n 0, 0, encode_var n (fst v1) 0) + else g) |> + (fun g' -> + if DFG.out_degree dfg v1 = 0 + 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, 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 = + List.fold_left (fun g n' -> + G.add_edge_e g (encode_bb n 1, -1, encode_bb n' 0) + ) constr succs + +module BFDFG = Graph.Path.BellmanFord(DFG)(struct + include DFG + type t = int + let weight = DFG.E.label + let compare = compare + let add = (+) + let zero = 0 + end) + +module TopoDFG = Graph.Topological.Make(DFG) + +let negate_graph constr = + DFG.fold_edges_e (function (v1, e, v2) -> fun g -> + DFG.add_edge_e g (v1, -e, v2) + ) constr DFG.empty + +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 + |> 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 maxf))) - 1), + encode_var n (fst v') 0) + ) constr (DFG.fold_vertex (fun v l -> + 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 = + | Mem + | SDiv + | UDiv + +type resources = { + res_mem: DFG.V.t list; + res_udiv: DFG.V.t list; + res_sdiv: DFG.V.t list; +} + +let find_resource = function + | RBload _ -> Some Mem + | RBstore _ -> Some Mem + | RBop (_, op, _, _) -> + ( match op with + | Odiv -> Some SDiv + | Odivu -> Some UDiv + | Omod -> Some SDiv + | Omodu -> Some UDiv + | _ -> None ) + | _ -> None + +let add_resource_constr n dfg constr = + let res = TopoDFG.fold (function (i, instr) -> + function {res_mem = ml; res_sdiv = sdl; res_udiv = udl} as r -> + match find_resource instr with + | Some SDiv -> {r with res_sdiv = (i, instr) :: sdl} + | Some UDiv -> {r with res_udiv = (i, instr) :: udl} + | Some Mem -> {r with res_mem = (i, instr) :: ml} + | None -> r + ) dfg {res_mem = []; res_sdiv = []; res_udiv = []} + in + let get_constraints l g = List.fold_left (fun gv v' -> + match gv with + | (g, None) -> (g, Some v') + | (g, Some v) -> + (G.add_edge_e g (encode_var n (fst v) 0, -1, encode_var n (fst v') 0), Some v') + ) (g, None) l |> fst + in + get_constraints (List.rev res.res_mem) constr + |> get_constraints (List.rev res.res_udiv) + |> get_constraints (List.rev res.res_sdiv) + +let gather_cfg_constraints c constr curr = + let (n, dfg) = curr in + match PTree.get (P.of_int n) c with + | 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 + |> List.filter (fun n' -> n' < n)) + |> add_cycle_constr 8 n dfg + |> add_resource_constr n dfg + +let rec intersperse s = function + | [] -> [] + | [ a ] -> [ a ] + | x :: xs -> x :: s :: intersperse s xs + +let print_objective constr = + let vars = G.fold_vertex (fun v1 l -> + match v1 with + | { sv_type = VarType _; sv_num = 0 } -> print_sv v1 :: l + | _ -> l + ) constr [] + in + "min: " ^ String.concat "" (intersperse " + " vars) ^ ";\n" + +let print_lp constr = + print_objective constr ^ + (G.fold_edges_e (function (e1, w, e2) -> fun s -> + s ^ sprintf "%s - %s <= %d;\n" (print_sv e1) (print_sv e2) w + ) constr "" |> + G.fold_vertex (fun v1 s -> + s ^ sprintf "int %s;\n" (print_sv v1) + ) constr) + +let update_schedule v = function Some l -> Some (v :: l) | None -> Some [ v ] + +let parse_soln (tree, bbtree) s = + let r = Str.regexp "var\\([0-9]+\\)n\\([0-9]+\\)_0[ ]+\\([0-9]+\\)" in + 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 (fn, oc) = Filename.open_temp_file "vericert_" "_lp_solve" in + fprintf oc "%s\n" (print_lp constr); + close_out oc; + + let res = Str.split (Str.regexp_string "\n") (read_process ("lp_solve " ^ fn)) + |> drop 3 + |> List.fold_left parse_soln (IMap.empty, 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 + List.fold_left (fun g v -> + List.fold_left (fun g' v' -> + let edges = DFG.find_all_edges dfg v v' in + List.fold_left DFG.add_edge_e g' edges + ) g l + ) dfg' l + +let rec all_successors dfg v = + List.concat (List.fold_left (fun l v -> + all_successors dfg v :: l + ) [] (DFG.succ dfg v)) + +let order_instr dfg = + DFG.fold_vertex (fun v li -> + if DFG.in_degree dfg v = 0 + then (List.map snd (v :: all_successors dfg v)) :: li + else li + ) dfg [] + +let combine_bb_schedule schedule s = + let i, st = s in + IMap.update st (update_schedule i) schedule + +(**let add_el dfg i l = + List.*) + +let check_in el = + List.exists (List.exists ((=) el)) + +let all_dfs dfg = + let roots = DFG.fold_vertex (fun v li -> + if DFG.in_degree dfg v = 0 then v :: li else li + ) dfg [] in + let dfg' = DFG.fold_edges (fun v1 v2 g -> DFG.add_edge g v2 v1) dfg dfg in + List.fold_left (fun a el -> + if check_in el a then a else + (DFGDfs.fold_component (fun v l -> v :: l) [] dfg' el) :: a) [] roots + +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 = + 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 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 + |> 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 -> + all_dfs x + |> List.map (subgraph x) + |> List.map (fun y -> + TopoDFG.fold (fun i l -> snd i :: l) y [] + |> List.rev))) body2 + (*|> (fun x -> TopoDFG.fold (fun i l -> snd i :: l) x []) + |> List.rev) body2*) + in + { bb_body = final_body2; + bb_exit = ctrl_flow + } + in + PTree.map f c + +let schedule entry (c : RTLBlock.bb RTLBlockInstr.code) = + let debug = true in + let transf_graph (_, dfg, _) = dfg in + let c' = PTree.map1 (fun x -> gather_bb_constraints false x |> transf_graph) c in + (*let _ = if debug then PTree.map (fun r o -> printf "##### %d #####\n%a\n\n" (P.to_int r) print_dfg o) c' else PTree.empty in*) + let cgraph = PTree.elements c' + |> List.map (function (x, y) -> (P.to_int x, y)) + |> List.fold_left (gather_cfg_constraints c) G.empty + in + let graph = open_out "constr_graph.dot" in + fprintf graph "%a\n" GDot.output_graph cgraph; + close_out graph; + let schedule' = solve_constraints cgraph in + (**IMap.iter (fun a b -> printf "##### %d #####\n%a\n\n" a (print_list print_tuple) b) schedule';*) + (**printf "Schedule: %a\n" (fun a x -> IMap.iter (fun d -> fprintf a "%d: %a\n" d (print_list print_tuple)) x) schedule';*) + transf_rtlpar c c' schedule' + +let rec find_reachable_states c e = + match PTree.get e c with + | Some { bb_exit = ex; _ } -> + e :: List.fold_left (fun x a -> List.concat [x; find_reachable_states c a]) [] + (successors_instr ex |> List.filter (fun x -> P.lt x e)) + | None -> assert false + +let add_to_tree c nt i = + match PTree.get i c with + | Some p -> PTree.set i p nt + | None -> assert false + +let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = + let scheduled = schedule f.fn_entrypoint f.fn_code in + let reachable = find_reachable_states scheduled f.fn_entrypoint + |> List.to_seq |> SS.of_seq |> SS.to_seq |> List.of_seq in + { fn_sig = f.fn_sig; + fn_params = f.fn_params; + fn_stacksize = f.fn_stacksize; + fn_code = scheduled (*List.fold_left (add_to_tree scheduled) PTree.empty reachable*); + fn_entrypoint = f.fn_entrypoint + } +#+end_src + +* RTLPargen +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargen.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpargen-main +#+begin_src coq +Require Import compcert.backend.Registers. +Require Import compcert.common.AST. +Require Import compcert.common.Globalenvs. +Require Import compcert.common.Memory. +Require Import compcert.common.Values. +Require Import compcert.lib.Floats. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. +Require compcert.verilog.Op. + +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLPar. +Require Import vericert.hls.RTLBlockInstr. +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. + +(** ** 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) + end. + +Fixpoint replicate {A} (n: nat) (l: A) := + match n with + | O => nil + | S n => l :: replicate n l + end. + +Definition merge''' x y := + match x, y with + | Some p1, Some p2 => Some (Pand p1 p2) + | Some p, None | None, Some p => Some p + | None, None => None + end. + +Definition merge'' x := + match x with + | ((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 := + match pa, pf with + | (p, a), (p', f) => (merge''' p p', f a) + end. + +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. + +(*map (fun x => (fst x, Econs (snd x) Enil)) pel*) +Definition merge' (pel: pred_expr) (tpel: predicated expression_list) := + predicated_map (uncurry Econs) (predicated_prod pel tpel). + +Fixpoint merge (pel: list pred_expr): predicated expression_list := + match pel with + | nil => NE.singleton (T, Enil) + | a :: b => merge' a (merge b) + end. + +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 := + 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 := + 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 := + NE.map (fun x => (fst x, (snd x) pa pb pc)) pf. + +Definition predicated_from_opt {A: Type} (p: option pred_op) (a: A) := + match p with + | Some p' => NE.singleton (p', a) + | None => NE.singleton (T, a) + end. + +#[local] Open Scope non_empty_scope. +#[local] Open Scope pred_op. + +Fixpoint NEfold_left {A B} (f: A -> B -> A) (l: NE.non_empty B) (a: A) : A := + match l with + | NE.singleton a' => f a a' + | a' ::| b => NEfold_left f b (f a a') + end. + +Fixpoint NEapp {A} (l m: NE.non_empty A) := + match l with + | NE.singleton a => a ::| m + | a ::| b => a ::| NEapp b m + end. + +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. + +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 + +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 + | RBop p op rl r => + f # (Reg r) <- + (app_predicated p + (f # (Reg r)) + (map_predicated (pred_ret (Eop op)) (merge (list_translation rl f)))) + | RBload p chunk addr rl r => + f # (Reg r) <- + (app_predicated p + (f # (Reg r)) + (map_predicated + (map_predicated (pred_ret (Eload chunk addr)) (merge (list_translation rl f))) + (f # Mem))) + | RBstore p chunk addr rl r => + f # Mem <- + (app_predicated p + (f # Mem) + (map_predicated + (map_predicated + (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)))) + 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. + +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. + +(** 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. *) + +Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := + match bb with + | nil => + match bbt with + | nil => true + | _ => false + end + | _ => true + end. + +Definition schedule_oracle (bb: RTLBlock.bblock) (bbt: RTLPar.bblock) : bool := + check (abstract_sequence empty (bb_body bb)) + (abstract_sequence empty (concat (concat (bb_body bbt)))) && + check_control_flow_instr (bb_exit bb) (bb_exit bbt) && + empty_trees (bb_body bb) (bb_body bbt). + +Definition check_scheduled_trees := beq2 schedule_oracle. + +Ltac solve_scheduled_trees_correct := + intros; unfold check_scheduled_trees in *; + match goal with + | [ H: context[beq2 _ _ _], x: positive |- _ ] => + rewrite beq2_correct in H; specialize (H x) + end; repeat destruct_match; crush. + +Lemma check_scheduled_trees_correct: + forall f1 f2 x y1, + check_scheduled_trees f1 f2 = true -> + PTree.get x f1 = Some y1 -> + exists y2, PTree.get x f2 = Some y2 /\ schedule_oracle y1 y2 = true. +Proof. solve_scheduled_trees_correct; eexists; crush. Qed. + +Lemma check_scheduled_trees_correct2: + forall f1 f2 x, + check_scheduled_trees f1 f2 = true -> + PTree.get x f1 = None -> + PTree.get x f2 = None. +Proof. solve_scheduled_trees_correct. Qed. + +(** ** Top-level Functions *) + +Parameter schedule : RTLBlock.function -> 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) + f.(fn_params) + f.(fn_stacksize) + tfcode + f.(fn_entrypoint)) + else + 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. +#+end_src +* RTLPargenproof +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargenproof.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpargenproof-imports +#+begin_src coq +Require Import compcert.backend.Registers. +Require Import compcert.common.AST. +Require Import compcert.common.Errors. +Require Import compcert.common.Linking. +Require Import compcert.common.Globalenvs. +Require Import compcert.common.Memory. +Require Import compcert.common.Values. +Require Import compcert.lib.Maps. + +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLPar. +Require Import vericert.hls.RTLBlockInstr. +Require Import vericert.hls.RTLPargen. +Require Import vericert.hls.Predicate. +Require Import vericert.hls.Abstr. + +#[local] Open Scope positive. +#[local] Open Scope forest. +#[local] Open Scope pred_op. +#+end_src + +** RTLBlock to abstract translation + +Correctness of translation from RTLBlock to the abstract interpretation language. + +#+name: rtlpargenproof-main +#+begin_src coq +(*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. + +Inductive state_lessdef : instr_state -> instr_state -> Prop := + state_lessdef_intro : + forall rs1 rs2 m1, + (forall x, rs1 !! x = rs2 !! x) -> + state_lessdef (mk_instr_state rs1 m1) (mk_instr_state rs2 m1). + +Ltac inv_simp := + repeat match goal with + | H: exists _, _ |- _ => inv H + end; simplify. + +*) + +Definition check_dest i r' := + match i with + | RBop p op rl r => (r =? r')%positive + | RBload p chunk addr rl r => (r =? r')%positive + | _ => false + end. + +Lemma check_dest_dec i r : {check_dest i r = true} + {check_dest i r = false}. +Proof. destruct (check_dest i r); tauto. Qed. + +Fixpoint check_dest_l l r := + match l with + | nil => false + | a :: b => check_dest a r || check_dest_l b r + end. + +Lemma check_dest_l_forall : + forall l r, + check_dest_l l r = false -> + Forall (fun x => check_dest x r = false) l. +Proof. induction l; crush. Qed. + +Lemma check_dest_l_dec i r : {check_dest_l i r = true} + {check_dest_l i r = false}. +Proof. destruct (check_dest_l i r); tauto. Qed. + +Lemma check_dest_update : + forall f i r, + check_dest i r = false -> + (update f i) # (Reg r) = f # (Reg r). +Proof. + destruct i; crush; try apply Pos.eqb_neq in H; apply genmap1; crush. +Qed. + +Lemma check_dest_l_forall2 : + forall l r, + Forall (fun x => check_dest x r = false) l -> + check_dest_l l r = false. +Proof. + induction l; crush. + inv H. apply orb_false_intro; crush. +Qed. + +Lemma check_dest_l_ex2 : + forall l r, + (exists a, In a l /\ check_dest a r = true) -> + check_dest_l l r = true. +Proof. + induction l; crush. + specialize (IHl r). inv H. + apply orb_true_intro; crush. + apply orb_true_intro; crush. + right. apply IHl. exists x. auto. +Qed. + +Lemma check_list_l_false : + forall l x r, + check_dest_l (l ++ x :: nil) r = false -> + check_dest_l l r = false /\ check_dest x r = false. +Proof. + simplify. + apply check_dest_l_forall in H. apply Forall_app in H. + simplify. apply check_dest_l_forall2; auto. + apply check_dest_l_forall in H. apply Forall_app in H. + simplify. inv H1. auto. +Qed. + +Lemma check_dest_l_ex : + forall l r, + check_dest_l l r = true -> + exists a, In a l /\ check_dest a r = true. +Proof. + induction l; crush. + destruct (check_dest a r) eqn:?; try solve [econstructor; crush]. + simplify. + exploit IHl. apply H. simplify. econstructor. simplify. right. eassumption. + auto. +Qed. + +Lemma check_list_l_true : + forall l x r, + check_dest_l (l ++ x :: nil) r = true -> + check_dest_l l r = true \/ check_dest x r = true. +Proof. + simplify. + apply check_dest_l_ex in H; simplify. + apply in_app_or in H. inv H. left. + apply check_dest_l_ex2. exists x0. auto. + inv H0; auto. +Qed. + +Lemma check_dest_l_dec2 l r : + {Forall (fun x => check_dest x r = false) l} + + {exists a, In a l /\ check_dest a r = true}. +Proof. + destruct (check_dest_l_dec l r); [right | left]; + auto using check_dest_l_ex, check_dest_l_forall. +Qed. + +Lemma abstr_comp : + forall l i f x x0, + abstract_sequence f (l ++ i :: nil) = x -> + abstract_sequence f l = x0 -> + x = update x0 i. +Proof. induction l; intros; crush; eapply IHl; eauto. Qed. + +(* + +Lemma gen_list_base: + forall FF ge sp l rs exps st1, + (forall x, @sem_value FF ge sp st1 (exps # (Reg x)) (rs !! x)) -> + sem_val_list ge sp st1 (list_translation l exps) rs ## l. +Proof. + induction l. + intros. simpl. constructor. + intros. simpl. eapply Scons; eauto. +Qed. + +Lemma check_dest_update2 : + forall f r rl op p, + (update f (RBop p op rl r)) # (Reg r) = Eop op (list_translation rl f) (f # Mem). +Proof. crush; rewrite map2; auto. Qed. + +Lemma check_dest_update3 : + forall f r rl p addr chunk, + (update f (RBload p chunk addr rl r)) # (Reg r) = Eload chunk addr (list_translation rl f) (f # Mem). +Proof. crush; rewrite map2; auto. Qed. + +Lemma abstract_seq_correct_aux: + forall FF ge sp i st1 st2 st3 f, + @step_instr FF ge sp st3 i st2 -> + sem ge sp st1 f st3 -> + sem ge sp st1 (update f i) st2. +Proof. + intros; inv H; simplify. + { simplify; eauto. } (*apply match_states_refl. }*) + { inv H0. inv H6. destruct st1. econstructor. simplify. + constructor. intros. + destruct (resource_eq (Reg res) (Reg x)). inv e. + rewrite map2. econstructor. eassumption. apply gen_list_base; eauto. + rewrite Regmap.gss. eauto. + assert (res <> x). { unfold not in *. intros. apply n. rewrite H0. auto. } + rewrite Regmap.gso by auto. + rewrite genmap1 by auto. auto. + + rewrite genmap1; crush. } + { inv H0. inv H7. constructor. constructor. intros. + destruct (Pos.eq_dec dst x); subst. + rewrite map2. econstructor; eauto. + apply gen_list_base. auto. rewrite Regmap.gss. auto. + rewrite genmap1. rewrite Regmap.gso by auto. auto. + unfold not in *; intros. inv H0. auto. + rewrite genmap1; crush. + } + { inv H0. inv H7. constructor. constructor; intros. + rewrite genmap1; crush. + rewrite map2. econstructor; eauto. + apply gen_list_base; auto. + } +Qed. + +Lemma regmap_list_equiv : + forall A (rs1: Regmap.t A) rs2, + (forall x, rs1 !! x = rs2 !! x) -> + forall rl, rs1##rl = rs2##rl. +Proof. induction rl; crush. Qed. + +Lemma sem_update_Op : + forall A ge sp st f st' r l o0 o m rs v, + @sem A ge sp st f st' -> + Op.eval_operation ge sp o0 rs ## l m = Some v -> + match_states st' (mk_instr_state rs m) -> + exists tst, + sem ge sp st (update f (RBop o o0 l r)) tst /\ match_states (mk_instr_state (Regmap.set r v rs) m) tst. +Proof. + intros. inv H1. simplify. + destruct st. + econstructor. simplify. + { constructor. + { constructor. intros. destruct (Pos.eq_dec x r); subst. + { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. + { inv H9. eapply gen_list_base; eauto. } + { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } + { inv H. rewrite genmap1; crush. eauto. } } + { constructor; eauto. intros. + destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed. + +Lemma sem_update_load : + forall A ge sp st f st' r o m a l m0 rs v a0, + @sem A ge sp st f st' -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.loadv m m0 a0 = Some v -> + match_states st' (mk_instr_state rs m0) -> + exists tst : instr_state, + sem ge sp st (update f (RBload o m a l r)) tst + /\ match_states (mk_instr_state (Regmap.set r v rs) m0) tst. +Proof. + intros. inv H2. pose proof H. inv H. inv H9. + destruct st. + econstructor; simplify. + { constructor. + { constructor. intros. + destruct (Pos.eq_dec x r); subst. + { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. + rewrite <- H6. eauto. + instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. auto. } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. } } + { rewrite genmap1; crush. eauto. } } + { constructor; auto; intros. destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed. + +Lemma sem_update_store : + forall A ge sp a0 m a l r o f st m' rs m0 st', + @sem A ge sp st f st' -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.storev m m0 a0 rs !! r = Some m' -> + match_states st' (mk_instr_state rs m0) -> + exists tst, sem ge sp st (update f (RBstore o m a l r)) tst + /\ match_states (mk_instr_state rs m') tst. +Proof. + intros. inv H2. pose proof H. inv H. inv H9. + destruct st. + econstructor; simplify. + { econstructor. + { econstructor; intros. rewrite genmap1; crush. } + { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. rewrite <- H6. + eauto. specialize (H6 r). rewrite H6. eauto. } } + { econstructor; eauto. } +Qed. + +Lemma sem_update : + forall A ge sp st x st' st'' st''' f, + sem ge sp st f st' -> + match_states st' st''' -> + @step_instr A ge sp st''' x st'' -> + exists tst, sem ge sp st (update f x) tst /\ match_states st'' tst. +Proof. + intros. destruct x; inv H1. + { econstructor. split. + apply sem_update_RBnop. eassumption. + apply match_states_commut. auto. } + { eapply sem_update_Op; eauto. } + { eapply sem_update_load; eauto. } + { eapply sem_update_store; eauto. } +Qed. + +Lemma sem_update2_Op : + forall A ge sp st f r l o0 o m rs v, + @sem A ge sp st f (mk_instr_state rs m) -> + Op.eval_operation ge sp o0 rs ## l m = Some v -> + sem ge sp st (update f (RBop o o0 l r)) (mk_instr_state (Regmap.set r v rs) m). +Proof. + intros. destruct st. constructor. + inv H. inv H6. + { constructor; intros. simplify. + destruct (Pos.eq_dec r x); subst. + { rewrite map2. econstructor. eauto. + apply gen_list_base. eauto. + rewrite Regmap.gss. auto. } + { rewrite genmap1; crush. rewrite Regmap.gso; auto. } } + { simplify. rewrite genmap1; crush. inv H. eauto. } +Qed. + +Lemma sem_update2_load : + forall A ge sp st f r o m a l m0 rs v a0, + @sem A ge sp st f (mk_instr_state rs m0) -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.loadv m m0 a0 = Some v -> + sem ge sp st (update f (RBload o m a l r)) (mk_instr_state (Regmap.set r v rs) m0). +Proof. + intros. simplify. inv H. inv H7. constructor. + { constructor; intros. destruct (Pos.eq_dec r x); subst. + { rewrite map2. rewrite Regmap.gss. econstructor; eauto. + apply gen_list_base; eauto. } + { rewrite genmap1; crush. rewrite Regmap.gso; eauto. } + } + { simplify. rewrite genmap1; crush. } +Qed. + +Lemma sem_update2_store : + forall A ge sp a0 m a l r o f st m' rs m0, + @sem A ge sp st f (mk_instr_state rs m0) -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.storev m m0 a0 rs !! r = Some m' -> + sem ge sp st (update f (RBstore o m a l r)) (mk_instr_state rs m'). +Proof. + intros. simplify. inv H. inv H7. constructor; simplify. + { econstructor; intros. rewrite genmap1; crush. } + { rewrite map2. econstructor; eauto. apply gen_list_base; eauto. } +Qed. + +Lemma sem_update2 : + forall A ge sp st x st' st'' f, + sem ge sp st f st' -> + @step_instr A ge sp st' x st'' -> + sem ge sp st (update f x) st''. +Proof. + intros. + destruct x; inv H0; + eauto using sem_update_RBnop, sem_update2_Op, sem_update2_load, sem_update2_store. +Qed. + +Lemma abstr_sem_val_mem : + forall A B ge tge st tst sp a, + ge_preserved ge tge -> + forall v m, + (@sem_mem A ge sp st a m /\ match_states st tst -> @sem_mem B tge sp tst a m) /\ + (@sem_value A ge sp st a v /\ match_states st tst -> @sem_value B tge sp tst a v). +Proof. + intros * H. + apply expression_ind2 with + + (P := fun (e1: expression) => + forall v m, + (@sem_mem A ge sp st e1 m /\ match_states st tst -> @sem_mem B tge sp tst e1 m) /\ + (@sem_value A ge sp st e1 v /\ match_states st tst -> @sem_value B tge sp tst e1 v)) + + (P0 := fun (e1: expression_list) => + forall lv, @sem_val_list A ge sp st e1 lv /\ match_states st tst -> @sem_val_list B tge sp tst e1 lv); + simplify; intros; simplify. + { inv H1. inv H2. constructor. } + { inv H2. inv H1. rewrite H0. constructor. } + { inv H3. } + { inv H3. inv H4. econstructor. apply H1; auto. simplify. eauto. constructor. auto. auto. + apply H0; simplify; eauto. constructor; eauto. + unfold ge_preserved in *. simplify. rewrite <- H2. auto. + } + { inv H3. } + { inv H3. inv H4. econstructor. apply H1; eauto; simplify; eauto. constructor; eauto. + apply H0; simplify; eauto. constructor; eauto. + inv H. rewrite <- H4. eauto. + auto. + } + { inv H4. inv H5. econstructor. apply H0; eauto. simplify; eauto. constructor; eauto. + apply H2; eauto. simplify; eauto. constructor; eauto. + apply H1; eauto. simplify; eauto. constructor; eauto. + inv H. rewrite <- H5. eauto. auto. + } + { inv H4. } + { inv H1. constructor. } + { inv H3. constructor; auto. apply H0; eauto. apply Mem.empty. } +Qed. + +Lemma abstr_sem_value : + forall a A B ge tge sp st tst v, + @sem_value A ge sp st a v -> + ge_preserved ge tge -> + match_states st tst -> + @sem_value B tge sp tst a v. +Proof. intros; eapply abstr_sem_val_mem; eauto; apply Mem.empty. Qed. + +Lemma abstr_sem_mem : + forall a A B ge tge sp st tst v, + @sem_mem A ge sp st a v -> + ge_preserved ge tge -> + match_states st tst -> + @sem_mem B tge sp tst a v. +Proof. intros; eapply abstr_sem_val_mem; eauto. Qed. + +Lemma abstr_sem_regset : + forall a a' A B ge tge sp st tst rs, + @sem_regset A ge sp st a rs -> + ge_preserved ge tge -> + (forall x, a # x = a' # x) -> + match_states st tst -> + exists rs', @sem_regset B tge sp tst a' rs' /\ (forall x, rs !! x = rs' !! x). +Proof. + inversion 1; intros. + inv H7. + econstructor. simplify. econstructor. intros. + eapply abstr_sem_value; eauto. rewrite <- H6. + eapply H0. constructor; eauto. + auto. +Qed. + +Lemma abstr_sem : + forall a a' A B ge tge sp st tst st', + @sem A ge sp st a st' -> + ge_preserved ge tge -> + (forall x, a # x = a' # x) -> + match_states st tst -> + exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. +Proof. + inversion 1; subst; intros. + inversion H4; subst. + exploit abstr_sem_regset; eauto; inv_simp. + do 3 econstructor; eauto. + rewrite <- H3. + eapply abstr_sem_mem; eauto. +Qed. + +Lemma abstract_execution_correct': + forall A B ge tge sp st' a a' st tst, + @sem A ge sp st a st' -> + ge_preserved ge tge -> + check a a' = true -> + match_states st tst -> + exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. +Proof. + intros; + pose proof (check_correct a a' H1); + eapply abstr_sem; eauto. +Qed. + +Lemma states_match : + forall st1 st2 st3 st4, + match_states st1 st2 -> + match_states st2 st3 -> + match_states st3 st4 -> + match_states st1 st4. +Proof. + intros * H1 H2 H3; destruct st1; destruct st2; destruct st3; destruct st4. + inv H1. inv H2. inv H3; constructor. + unfold regs_lessdef in *. intros. + repeat match goal with + | H: forall _, _, r : positive |- _ => specialize (H r) + end. + congruence. + auto. +Qed. + +Lemma step_instr_block_same : + forall ge sp st st', + step_instr_block ge sp st nil st' -> + st = st'. +Proof. inversion 1; auto. Qed. + +Lemma step_instr_seq_same : + forall ge sp st st', + step_instr_seq ge sp st nil st' -> + st = st'. +Proof. inversion 1; auto. Qed. + +Lemma sem_update' : + forall A ge sp st a x st', + sem ge sp st (update (abstract_sequence empty a) x) st' -> + exists st'', + @step_instr A ge sp st'' x st' /\ + sem ge sp st (abstract_sequence empty a) st''. +Proof. + Admitted. + +Lemma rtlpar_trans_correct : + forall bb ge sp sem_st' sem_st st, + sem ge sp sem_st (abstract_sequence empty (concat (concat bb))) sem_st' -> + match_states sem_st st -> + exists st', RTLPar.step_instr_block ge sp st bb st' + /\ match_states sem_st' st'. +Proof. + induction bb using rev_ind. + { repeat econstructor. eapply abstract_interp_empty3 in H. + inv H. inv H0. constructor; congruence. } + { simplify. inv H0. repeat rewrite concat_app in H. simplify. + rewrite app_nil_r in H. + exploit sem_separate; eauto; inv_simp. + repeat econstructor. admit. admit. + } +Admitted. + +(*Lemma abstract_execution_correct_ld: + forall bb bb' cfi ge tge sp st st' tst, + RTLBlock.step_instr_list ge sp st bb st' -> + ge_preserved ge tge -> + schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi) = true -> + match_states_ld st tst -> + exists tst', RTLPar.step_instr_block tge sp tst bb' tst' + /\ match_states st' tst'. +Proof. + intros.*) +*) + +Lemma match_states_list : + forall A (rs: Regmap.t A) rs', + (forall r, rs !! r = rs' !! r) -> + forall l, rs ## l = rs' ## l. +Proof. induction l; crush. Qed. + +Lemma PTree_matches : + forall A (v: A) res rs rs', + (forall r, rs !! r = rs' !! r) -> + forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x. +Proof. + intros; destruct (Pos.eq_dec x res); subst; + [ repeat rewrite Regmap.gss by auto + | repeat rewrite Regmap.gso by auto ]; auto. +Qed. + +Lemma abstract_interp_empty3 : + forall A ctx st', + @sem A ctx empty st' -> match_states (ctx_is ctx) st'. +Proof. + inversion 1; subst; simplify. destruct ctx. + destruct ctx_is. + constructor; intros. + - inv H0. specialize (H3 x). inv H3. inv H8. reflexivity. + - inv H1. specialize (H3 x). inv H3. inv H8. reflexivity. + - inv H2. inv H8. reflexivity. +Qed. + +Lemma step_instr_matches : + forall A a ge sp st st', + @step_instr A ge sp st a st' -> + forall tst, + match_states st tst -> + exists tst', step_instr ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction 1; simplify; + match goal with H: match_states _ _ |- _ => inv H end; + try solve [repeat econstructor; try erewrite match_states_list; + try apply PTree_matches; eauto; + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end]. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + apply PTree_matches; assumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + econstructor; auto. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + apply PTree_matches; assumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + econstructor; auto. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + erewrite <- eval_predf_pr_equiv; eassumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + erewrite <- eval_predf_pr_equiv; eassumption. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + - admit. Admitted. + +Lemma step_instr_list_matches : + forall a ge sp st st', + step_instr_list ge sp st a st' -> + forall tst, match_states st tst -> + exists tst', step_instr_list ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction a; intros; inv H; + try (exploit step_instr_matches; eauto; []; simplify; + exploit IHa; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma step_instr_seq_matches : + forall a ge sp st st', + step_instr_seq ge sp st a st' -> + forall tst, match_states st tst -> + exists tst', step_instr_seq ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction a; intros; inv H; + try (exploit step_instr_list_matches; eauto; []; simplify; + exploit IHa; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma step_instr_block_matches : + forall bb ge sp st st', + step_instr_block ge sp st bb st' -> + forall tst, match_states st tst -> + exists tst', step_instr_block ge sp tst bb tst' + /\ match_states st' tst'. +Proof. + induction bb; intros; inv H; + try (exploit step_instr_seq_matches; eauto; []; simplify; + exploit IHbb; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma rtlblock_trans_correct' : + forall bb ge sp st x st'', + RTLBlock.step_instr_list ge sp st (bb ++ x :: nil) st'' -> + exists st', RTLBlock.step_instr_list ge sp st bb st' + /\ step_instr ge sp st' x st''. +Proof. + induction bb. + crush. exists st. + split. constructor. inv H. inv H6. auto. + crush. inv H. exploit IHbb. eassumption. simplify. + econstructor. split. + econstructor; eauto. eauto. +Qed. + +Lemma abstract_interp_empty A st : @sem A st empty (ctx_is st). +Proof. destruct st, ctx_is. simpl. repeat econstructor. Qed. + +Lemma abstract_seq : + forall l f i, + abstract_sequence f (l ++ i :: nil) = update (abstract_sequence f l) i. +Proof. induction l; crush. Qed. + +Lemma abstract_sequence_update : + forall l r f, + check_dest_l l r = false -> + (abstract_sequence f l) # (Reg r) = f # (Reg r). +Proof. + induction l using rev_ind; crush. + rewrite abstract_seq. rewrite check_dest_update. apply IHl. + apply check_list_l_false in H. tauto. + apply check_list_l_false in H. tauto. +Qed. + +(*Lemma sem_separate : + forall A ctx b a st', + sem ctx (abstract_sequence empty (a ++ b)) st' -> + exists st'', + @sem A ctx (abstract_sequence empty a) st'' + /\ @sem A (mk_ctx st'' (ctx_sp ctx) (ctx_ge ctx)) (abstract_sequence empty b) st'. +Proof. + induction b using rev_ind; simplify. + { econstructor. simplify. rewrite app_nil_r in H. eauto. apply abstract_interp_empty. } + { simplify. rewrite app_assoc in H. rewrite abstract_seq in H. + exploit sem_update'; eauto; simplify. + exploit IHb; eauto; inv_simp. + econstructor; split; eauto. + rewrite abstract_seq. + eapply sem_update2; eauto. + } +Qed.*) + +Lemma sem_update_RBnop : + forall A ctx f st', + @sem A ctx f st' -> sem ctx (update f RBnop) st'. +Proof. auto. Qed. + +Lemma sem_update_Op : + forall A ge sp ist f st' r l o0 o m rs v ps, + @sem A (mk_ctx ist sp ge) f st' -> + eval_predf ps o = true -> + Op.eval_operation ge sp o0 (rs ## l) m = Some v -> + match_states st' (mk_instr_state rs ps m) -> + exists tst, + sem (mk_ctx ist sp ge) (update f (RBop (Some o) o0 l r)) tst + /\ match_states (mk_instr_state (Regmap.set r v rs) ps m) tst. +Proof. + intros. inv H1. inv H. inv H1. inv H3. simplify. + econstructor. simplify. + { constructor; try constructor; intros; try solve [rewrite genmap1; now eauto]. + destruct (Pos.eq_dec x r); subst. + { rewrite map2. specialize (H1 r). inv H1. +(*} + } + destruct st. + econstructor. simplify. + { constructor. + { constructor. intros. destruct (Pos.eq_dec x r); subst. + { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. + { inv H9. eapply gen_list_base; eauto. } + { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } + { inv H. rewrite genmap1; crush. eauto. } } + { constructor; eauto. intros. + destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed.*) Admitted. + +Lemma sem_update : + forall A ge sp st x st' st'' st''' f, + sem (mk_ctx st sp ge) f st' -> + 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. + 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', + RTLBlock.step_instr_list ge sp st bb st' -> + forall tst, + match_states st tst -> + exists tst', sem (mk_ctx tst sp ge) (abstract_sequence empty bb) tst' + /\ match_states st' tst'. +Proof. + induction bb using rev_ind; simplify. + { econstructor. simplify. apply abstract_interp_empty. + inv H. auto. } + { apply rtlblock_trans_correct' in H. simplify. + rewrite abstract_seq. + exploit IHbb; try eassumption; []; simplify. + exploit sem_update. apply H1. symmetry; eassumption. + eauto. simplify. econstructor. split. apply H3. + auto. } +Qed. + +Lemma abstract_execution_correct: + forall bb bb' cfi cfi' ge tge sp st st' tst, + RTLBlock.step_instr_list ge sp st bb st' -> + ge_preserved ge tge -> + schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi') = true -> + match_states st tst -> + exists tst', RTLPar.step_instr_block tge sp tst bb' tst' + /\ match_states st' tst'. +Proof. + intros. + unfold schedule_oracle in *. simplify. unfold empty_trees in H4. + exploit rtlblock_trans_correct; try eassumption; []; simplify. +(*) exploit abstract_execution_correct'; + try solve [eassumption | apply state_lessdef_match_sem; eassumption]. + apply match_states_commut. eauto. inv_simp. + exploit rtlpar_trans_correct; try eassumption; []; inv_simp. + exploit step_instr_block_matches; eauto. apply match_states_commut; eauto. inv_simp. + repeat match goal with | H: match_states _ _ |- _ => inv H end. + do 2 econstructor; eauto. + econstructor; congruence. +Qed.*)Admitted. + +Definition match_prog (prog : RTLBlock.program) (tprog : RTLPar.program) := + match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq prog tprog. + +Inductive match_stackframes: RTLBlock.stackframe -> RTLPar.stackframe -> Prop := +| match_stackframe: + forall f tf res sp pc rs rs' ps ps', + transl_function f = OK tf -> + (forall x, rs !! x = rs' !! x) -> + (forall x, ps !! x = ps' !! x) -> + match_stackframes (Stackframe res f sp pc rs ps) + (Stackframe res tf sp pc rs' ps'). + +Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop := +| match_state: + forall sf f sp pc rs rs' m sf' tf ps ps' + (TRANSL: transl_function f = OK tf) + (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_returnstate: + forall stack stack' v m + (STACKS: list_forall2 match_stackframes stack stack'), + match_states (Returnstate stack v m) + (Returnstate stack' v m) +| match_callstate: + forall stack stack' f tf args m + (TRANSL: transl_fundef f = OK tf) + (STACKS: list_forall2 match_stackframes stack stack'), + match_states (Callstate stack f args m) + (Callstate stack' tf args m). + +Section CORRECTNESS. + + Context (prog: RTLBlock.program) (tprog : RTLPar.program). + Context (TRANSL: match_prog prog tprog). + + Let ge : RTLBlock.genv := Globalenvs.Genv.globalenv prog. + Let tge : RTLPar.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. + Hint Resolve symbols_preserved : rtlgp. + + Lemma function_ptr_translated: + forall (b: Values.block) (f: RTLBlock.fundef), + Genv.find_funct_ptr ge b = Some f -> + exists tf, + Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = Errors.OK tf. + Proof using TRANSL. + intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. + intros (cu & tf & P & Q & R); exists tf; auto. + Qed. + + Lemma functions_translated: + forall (v: Values.val) (f: RTLBlock.fundef), + Genv.find_funct ge v = Some f -> + exists tf, + Genv.find_funct tge v = Some tf /\ transl_fundef f = Errors.OK tf. + Proof using TRANSL. + intros. exploit (Genv.find_funct_match TRANSL); eauto. + intros (cu & tf & P & Q & R); exists tf; auto. + Qed. + + Lemma senv_preserved: + Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). + Proof (Genv.senv_transf_partial TRANSL). + Hint Resolve senv_preserved : rtlgp. + + Lemma sig_transl_function: + forall (f: RTLBlock.fundef) (tf: RTLPar.fundef), + transl_fundef f = OK tf -> + funsig tf = funsig f. + Proof using . + unfold transl_fundef, transf_partial_fundef, transl_function; intros; + repeat destruct_match; crush; + match goal with H: OK _ = OK _ |- _ => inv H end; auto. + Qed. + Hint Resolve sig_transl_function : rtlgp. + + Hint Resolve Val.lessdef_same : rtlgp. + Hint Resolve regs_lessdef_regs : rtlgp. + + Lemma find_function_translated: + forall ros rs rs' f, + (forall x, rs !! x = rs' !! x) -> + find_function ge ros rs = Some f -> + exists tf, find_function tge ros rs' = Some tf + /\ transl_fundef f = OK tf. + Proof using TRANSL. + Ltac ffts := match goal with + | [ H: forall _, Val.lessdef _ _, r: Registers.reg |- _ ] => + specialize (H r); inv H + | [ H: Vundef = ?r, H1: Genv.find_funct _ ?r = Some _ |- _ ] => + rewrite <- H in H1 + | [ H: Genv.find_funct _ Vundef = Some _ |- _] => solve [inv H] + | _ => solve [exploit functions_translated; eauto] + end. + destruct ros; simplify; try rewrite <- H; + [| rewrite symbols_preserved; destruct_match; + try (apply function_ptr_translated); crush ]; + intros; + repeat ffts. + Qed. + + Lemma schedule_oracle_nil: + forall bb cfi, + schedule_oracle {| bb_body := nil; bb_exit := cfi |} bb = true -> + bb_body bb = nil /\ bb_exit bb = cfi. + Proof using . + unfold schedule_oracle, check_control_flow_instr. + simplify; repeat destruct_match; crush. + Qed. + + Lemma schedule_oracle_nil2: + forall cfi, + schedule_oracle {| bb_body := nil; bb_exit := cfi |} + {| bb_body := nil; bb_exit := cfi |} = true. + Proof using . + unfold schedule_oracle, check_control_flow_instr. + simplify; repeat destruct_match; crush. + Qed. + + Lemma eval_op_eq: + forall (sp0 : Values.val) (op : Op.operation) (vl : list Values.val) m, + Op.eval_operation ge sp0 op vl m = Op.eval_operation tge sp0 op vl m. + Proof using TRANSL. + intros. + destruct op; auto; unfold Op.eval_operation, Genv.symbol_address, Op.eval_addressing32; + [| destruct a; unfold Genv.symbol_address ]; + try rewrite symbols_preserved; auto. + Qed. + Hint Resolve eval_op_eq : rtlgp. + + Lemma eval_addressing_eq: + forall sp addr vl, + Op.eval_addressing ge sp addr vl = Op.eval_addressing tge sp addr vl. + Proof using TRANSL. + intros. + destruct addr; + unfold Op.eval_addressing, Op.eval_addressing32; + unfold Genv.symbol_address; + try rewrite symbols_preserved; auto. + Qed. + Hint Resolve eval_addressing_eq : rtlgp. + + Lemma ge_preserved_lem: + ge_preserved ge tge. + Proof using TRANSL. + unfold ge_preserved. + eauto with rtlgp. + Qed. + Hint Resolve ge_preserved_lem : rtlgp. + + Lemma lessdef_regmap_optget: + forall or rs rs', + regs_lessdef rs rs' -> + Val.lessdef (regmap_optget or Vundef rs) (regmap_optget or Vundef rs'). + Proof using. destruct or; crush. Qed. + Hint Resolve lessdef_regmap_optget : rtlgp. + + Lemma regmap_equiv_lessdef: + forall rs rs', + (forall x, rs !! x = rs' !! x) -> + regs_lessdef rs rs'. + Proof using. + intros; unfold regs_lessdef; intros. + rewrite H. apply Val.lessdef_refl. + Qed. + Hint Resolve regmap_equiv_lessdef : rtlgp. + + Lemma int_lessdef: + forall rs rs', + regs_lessdef rs rs' -> + (forall arg v, + rs !! arg = Vint v -> + rs' !! arg = Vint v). + Proof using. intros ? ? H; intros; specialize (H arg); inv H; crush. Qed. + Hint Resolve int_lessdef : rtlgp. + + Ltac semantics_simpl := + match goal with + | [ H: match_states _ _ |- _ ] => + let H2 := fresh "H" in + learn H as H2; inv H2 + | [ H: transl_function ?f = OK _ |- _ ] => + let H2 := fresh "TRANSL" in + learn H as H2; + unfold transl_function in H2; + destruct (check_scheduled_trees + (@fn_code RTLBlock.bb f) + (@fn_code RTLPar.bb (schedule f))) eqn:?; + [| discriminate ]; inv H2 + | [ H: context[check_scheduled_trees] |- _ ] => + let H2 := fresh "CHECK" in + learn H as H2; + eapply check_scheduled_trees_correct in H2; [| solve [eauto] ] + | [ H: schedule_oracle {| bb_body := nil; bb_exit := _ |} ?bb = true |- _ ] => + let H2 := fresh "SCHED" in + learn H as H2; + apply schedule_oracle_nil in H2 + | [ H: find_function _ _ _ = Some _, H2: forall x, ?rs !! x = ?rs' !! x |- _ ] => + learn H; exploit find_function_translated; try apply H2; eauto; inversion 1 + | [ H: Mem.free ?m _ _ _ = Some ?m', H2: Mem.extends ?m ?m'' |- _ ] => + learn H; exploit Mem.free_parallel_extends; eauto; intros + | [ H: Events.eval_builtin_args _ _ _ _ _ _, H2: regs_lessdef ?rs ?rs' |- _ ] => + let H3 := fresh "H" in + learn H; exploit Events.eval_builtin_args_lessdef; [apply H2 | | |]; + eauto with rtlgp; intro H3; learn H3 + | [ H: Events.external_call _ _ _ _ _ _ _ |- _ ] => + let H2 := fresh "H" in + learn H; exploit Events.external_call_mem_extends; + eauto; intro H2; learn H2 + | [ H: exists _, _ |- _ ] => inv H + | _ => progress simplify + end. + + Hint Resolve Events.eval_builtin_args_preserved : rtlgp. + Hint Resolve Events.external_call_symbols_preserved : rtlgp. + Hint Resolve set_res_lessdef : rtlgp. + Hint Resolve set_reg_lessdef : rtlgp. + Hint Resolve Op.eval_condition_lessdef : rtlgp. + + Hint Constructors Events.eval_builtin_arg: barg. + + Lemma eval_builtin_arg_eq: + forall A ge a v1 m1 e1 e2 sp, + (forall x, e1 x = e2 x) -> + @Events.eval_builtin_arg A ge e1 sp m1 a v1 -> + Events.eval_builtin_arg ge e2 sp m1 a v1. +Proof. induction 2; try rewrite H; eauto with barg. Qed. + + Lemma eval_builtin_args_eq: + forall A ge e1 sp m1 e2 al vl1, + (forall x, e1 x = e2 x) -> + @Events.eval_builtin_args A ge e1 sp m1 al vl1 -> + Events.eval_builtin_args ge e2 sp m1 al vl1. + Proof. + induction 2. + - econstructor; split. + - exploit eval_builtin_arg_eq; eauto. intros. + destruct IHlist_forall2 as [| y]. constructor; eauto. + constructor. constructor; auto. + constructor; eauto. + Qed. + + Lemma step_cf_instr_correct: + forall cfi t s s', + step_cf_instr ge s cfi t s' -> + forall r, + match_states s r -> + exists r', step_cf_instr tge r cfi t r' /\ match_states s' r'. + Proof using TRANSL. + induction 1; repeat semantics_simpl; + try solve [repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp)]. + { do 3 (try erewrite match_states_list by eauto; econstructor; eauto with rtlgp). + eapply eval_builtin_args_eq. eapply REG. + eapply Events.eval_builtin_args_preserved. eapply symbols_preserved. + eauto. + intros. + unfold regmap_setres. destruct res. + destruct (Pos.eq_dec x0 x); subst. + repeat rewrite Regmap.gss; auto. + repeat rewrite Regmap.gso; auto. + eapply REG. eapply REG. + } + { repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). + unfold regmap_optget. destruct or. rewrite REG. constructor; eauto. + constructor; eauto. + } + { exploit IHstep_cf_instr; eauto. simplify. + repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). + erewrite eval_predf_pr_equiv; eauto. + } + Qed. + + Theorem transl_step_correct : + forall (S1 : RTLBlock.state) t S2, + RTLBlock.step ge S1 t S2 -> + forall (R1 : RTLPar.state), + match_states S1 R1 -> + exists R2, Smallstep.plus RTLPar.step tge R1 t R2 /\ match_states S2 R2. + Proof. + + induction 1; repeat semantics_simpl. + + { destruct bb; destruct x. + assert (bb_exit = bb_exit0). + { unfold schedule_oracle in *. simplify. + unfold check_control_flow_instr in *. + destruct_match; crush. + } + subst. + + exploit abstract_execution_correct; try eassumption. eapply ge_preserved_lem. + econstructor; eauto. + simplify. destruct x. inv H7. + + exploit step_cf_instr_correct; try eassumption. econstructor; eauto. + simplify. + + econstructor. econstructor. eapply Smallstep.plus_one. econstructor. + 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. + 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. } + { inv STACKS. inv H2. repeat econstructor; eauto. + intros. apply PTree_matches; eauto. } + Qed. + + Lemma transl_initial_states: + forall S, + RTLBlock.initial_state prog S -> + exists R, RTLPar.initial_state tprog R /\ match_states S R. + Proof. + induction 1. + exploit function_ptr_translated; eauto. intros [tf [A B]]. + econstructor; split. + econstructor. apply (Genv.init_mem_transf_partial TRANSL); eauto. + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto. + symmetry; eapply match_program_main; eauto. + eexact A. + rewrite <- H2. apply sig_transl_function; auto. + constructor. auto. constructor. + Qed. + + Lemma transl_final_states: + forall S R r, + match_states S R -> RTLBlock.final_state S r -> RTLPar.final_state R r. + Proof. + intros. inv H0. inv H. inv STACKS. constructor. + Qed. + + Theorem transf_program_correct: + Smallstep.forward_simulation (RTLBlock.semantics prog) (RTLPar.semantics tprog). + Proof. + eapply Smallstep.forward_simulation_plus. + apply senv_preserved. + eexact transl_initial_states. + eexact transl_final_states. + exact transl_step_correct. + Qed. + +End CORRECTNESS. +#+end_src + +* License + +#+name: license +#+begin_src coq :tangle no +(* + * 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 . + *) +#+end_src diff --git a/lit/scheduling.org b/lit/scheduling.org deleted file mode 100644 index 728dfc5..0000000 --- a/lit/scheduling.org +++ /dev/null @@ -1,3303 +0,0 @@ -#+title: Scheduling -#+author: Yann Herklotz -#+email: yann@yannherklotz.com - -* Language Definitions - -** RTLBlockInstr -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockInstr.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblockinstr-imports -#+begin_src coq -Require Import Coq.micromega.Lia. - -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.Values. -Require Import compcert.lib.Integers. -Require Import compcert.lib.Maps. -Require Import compcert.verilog.Op. - -Require Import Predicate. -Require Import Vericertlib. -#+end_src - -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]). - -#+name: rtlblockinstr-instr-def -#+begin_src coq -Definition node := positive. - -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. -#+end_src - -*** Control-Flow Instruction Definition - -These are the instructions that count as control-flow, and will be placed at the end of the basic -blocks. - -#+name: rtlblockinstr-cf-instr-def -#+begin_src coq -Inductive cf_instr : Type := -| RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr -| RBtailcall : signature -> reg + ident -> list reg -> cf_instr -| RBbuiltin : external_function -> list (builtin_arg reg) -> - builtin_res reg -> node -> cf_instr -| RBcond : condition -> list reg -> node -> node -> cf_instr -| RBjumptable : reg -> list node -> cf_instr -| RBreturn : option reg -> cf_instr -| RBgoto : node -> cf_instr -| RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. -#+end_src - -*** Helper functions - -#+name: rtlblockinstr-helpers -#+begin_src coq -Fixpoint successors_instr (i : cf_instr) : list node := - match i with - | RBcall sig ros args res s => s :: nil - | RBtailcall sig ros args => nil - | RBbuiltin ef args res s => s :: nil - | RBcond cond args ifso ifnot => ifso :: ifnot :: nil - | RBjumptable arg tbl => tbl - | RBreturn optarg => nil - | RBgoto n => n :: nil - | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil) - end. - -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) - | RBload p chunk addr args dst => - fold_left Pos.max args (Pos.max dst m) - | RBstore p chunk addr args src => - fold_left Pos.max args (Pos.max src m) - | RBsetpred p' c args p => - 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)) - | RBcall sig (inr id) args res s => - fold_left Pos.max args (Pos.max res m) - | RBtailcall sig (inl r) args => - fold_left Pos.max args (Pos.max r m) - | RBtailcall sig (inr id) args => - 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) - | RBcond cond args ifso ifnot => fold_left Pos.max args m - | RBjumptable arg tbl => Pos.max arg m - | RBreturn None => m - | RBreturn (Some arg) => Pos.max arg m - | RBgoto n => m - | RBpred_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. -#+end_src - -**** 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. - -#+name: rtlblockinstr-instr-state -#+begin_src coq -Record instr_state := mk_instr_state { - is_rs: regset; - is_ps: predset; - is_mem: mem; -}. -#+end_src - -*** Top-Level Type Definitions - -#+name: rtlblockinstr-type-def -#+begin_src coq -Section DEFINITION. - - Context {bblock_body: Type}. - - 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_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. - -End DEFINITION. -#+end_src - -*** Semantics - -#+name: rtlblockinstr-semantics -#+begin_src coq -Section RELSEM. - - Context {bblock_body : Type}. - - Definition genv := Genv.t (@fundef bblock_body) unit. - - 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_RBnop: - 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 - | 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 - | 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 - | 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. - - 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, - 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') - 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) (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') - | 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) - | 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) - | 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) - 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) - | 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'. - -End RELSEM. -#+end_src - -** RTLBlock -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlock.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblock-main -#+begin_src coq -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.RTLBlockInstr. - -Definition bb := list instr. - -Definition bblock := @bblock bb. -Definition code := @code bb. -Definition function := @function bb. -Definition fundef := @fundef bb. -Definition program := @program bb. -Definition funsig := @funsig bb. -Definition stackframe := @stackframe 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). - - 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_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). -#+end_src - -** RTLPar -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPar.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpar-main -#+begin_src coq -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.RTLBlockInstr. - -Definition bb := list (list (list instr)). - -Definition bblock := @bblock bb. -Definition code := @code bb. -Definition function := @function bb. -Definition fundef := @fundef bb. -Definition program := @program bb. -Definition funsig := @funsig bb. -Definition stackframe := @stackframe bb. -Definition state := @state bb. -Definition genv := @genv bb. - -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_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 -> bb -> 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: 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 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') - | 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. -#+end_src -* Scheduler -:PROPERTIES: -:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Schedule.ml -:END: - -#+begin_src ocaml :comments no :padline no :exports none -<> -#+end_src - -#+name: scheduler-main -#+begin_src ocaml -open Printf -open Clflags -open Camlcoq -open Datatypes -open Coqlib -open Maps -open AST -open Kildall -open Op -open RTLBlockInstr -open Predicate -open RTLBlock -open HTL -open Verilog -open HTLgen -open HTLMonad -open HTLMonadExtra - -module SS = Set.Make(P) - -type svtype = - | BBType of int - | VarType of int * int - -type sv = { - sv_type: svtype; - sv_num: int; -} - -let print_sv v = - match v with - | { sv_type = BBType bbi; sv_num = n } -> sprintf "bb%d_%d" bbi n - | { sv_type = VarType (bbi, i); sv_num = n } -> sprintf "var%dn%d_%d" bbi i n - -module G = Graph.Persistent.Digraph.ConcreteLabeled(struct - type t = sv - let compare = compare - let equal = (=) - let hash = Hashtbl.hash -end)(struct - type t = int - let compare = compare - let hash = Hashtbl.hash - let equal = (=) - let default = 0 -end) - -module GDot = Graph.Graphviz.Dot(struct - let graph_attributes _ = [] - let default_vertex_attributes _ = [] - let vertex_name = print_sv - let vertex_attributes _ = [] - let get_subgraph _ = None - let default_edge_attributes _ = [] - let edge_attributes _ = [] - - include G - end) - -module DFG = Graph.Persistent.Digraph.ConcreteLabeled(struct - type t = int * instr - let compare = compare - let equal = (=) - let hash = Hashtbl.hash -end)(struct - type t = int - let compare = compare - let hash = Hashtbl.hash - let equal = (=) - let default = 0 -end) - -module DFGSimp = Graph.Persistent.Graph.Concrete(struct - type t = int * instr - let compare = compare - let equal = (=) - let hash = Hashtbl.hash - end) - -let convert dfg = - DFG.fold_vertex (fun v g -> DFGSimp.add_vertex g v) dfg DFGSimp.empty - |> DFG.fold_edges (fun v1 v2 g -> DFGSimp.add_edge (DFGSimp.add_edge g v1 v2) v2 v1) dfg - -let reg r = sprintf "r%d" (P.to_int r) -let print_pred r = sprintf "p%d" (P.to_int r) - -let print_instr = function - | RBnop -> "" - | RBload (_, _, _, _, r) -> sprintf "load(%s)" (reg r) - | RBstore (_, _, _, _, r) -> sprintf "store(%s)" (reg r) - | RBsetpred (_, _, _, p) -> sprintf "setpred(%s)" (print_pred p) - | RBop (_, op, args, d) -> - (match op, args with - | Omove, _ -> "mov" - | Ointconst n, _ -> sprintf "%s=%ld" (reg d) (camlint_of_coqint n) - | Olongconst n, _ -> sprintf "%s=%LdL" (reg d) (camlint64_of_coqint n) - | Ofloatconst n, _ -> sprintf "%s=%.15F" (reg d) (camlfloat_of_coqfloat n) - | Osingleconst n, _ -> sprintf "%s=%.15Ff" (reg d) (camlfloat_of_coqfloat32 n) - | Oindirectsymbol id, _ -> sprintf "%s=&%s" (reg d) (extern_atom id) - | Ocast8signed, [r1] -> sprintf "%s=int8signed(%s)" (reg d) (reg r1) - | Ocast8unsigned, [r1] -> sprintf "%s=int8unsigned(%s)" (reg d) (reg r1) - | Ocast16signed, [r1] -> sprintf "%s=int16signed(%s)" (reg d) (reg r1) - | Ocast16unsigned, [r1] -> sprintf "%s=int16unsigned(%s)" (reg d) (reg r1) - | Oneg, [r1] -> sprintf "%s=-%s" (reg d) (reg r1) - | Osub, [r1;r2] -> sprintf "%s=%s-%s" (reg d) (reg r1) (reg r2) - | Omul, [r1;r2] -> sprintf "%s=%s*%s" (reg d) (reg r1) (reg r2) - | Omulimm n, [r1] -> sprintf "%s=%s*%ld" (reg d) (reg r1) (camlint_of_coqint n) - | Omulhs, [r1;r2] -> sprintf "%s=mulhs(%s,%s)" (reg d) (reg r1) (reg r2) - | Omulhu, [r1;r2] -> sprintf "%s=mulhu(%s,%s)" (reg d) (reg r1) (reg r2) - | Odiv, [r1;r2] -> sprintf "%s=%s /s %s" (reg d) (reg r1) (reg r2) - | Odivu, [r1;r2] -> sprintf "%s=%s /u %s" (reg d) (reg r1) (reg r2) - | Omod, [r1;r2] -> sprintf "%s=%s %%s %s" (reg d) (reg r1) (reg r2) - | Omodu, [r1;r2] -> sprintf "%s=%s %%u %s" (reg d) (reg r1) (reg r2) - | Oand, [r1;r2] -> sprintf "%s=%s & %s" (reg d) (reg r1) (reg r2) - | Oandimm n, [r1] -> sprintf "%s=%s & %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oor, [r1;r2] -> sprintf "%s=%s | %s" (reg d) (reg r1) (reg r2) - | Oorimm n, [r1] -> sprintf "%s=%s | %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oxor, [r1;r2] -> sprintf "%s=%s ^ %s" (reg d) (reg r1) (reg r2) - | Oxorimm n, [r1] -> sprintf "%s=%s ^ %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Onot, [r1] -> sprintf "%s=not(%s)" (reg d) (reg r1) - | Oshl, [r1;r2] -> sprintf "%s=%s << %s" (reg d) (reg r1) (reg r2) - | Oshlimm n, [r1] -> sprintf "%s=%s << %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshr, [r1;r2] -> sprintf "%s=%s >>s %s" (reg d) (reg r1) (reg r2) - | Oshrimm n, [r1] -> sprintf "%s=%s >>s %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrximm n, [r1] -> sprintf "%s=%s >>x %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshru, [r1;r2] -> sprintf "%s=%s >>u %s" (reg d) (reg r1) (reg r2) - | Oshruimm n, [r1] -> sprintf "%s=%s >>u %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Ororimm n, [r1] -> sprintf "%s=%s ror %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshldimm n, [r1;r2] -> sprintf "%s=(%s, %s) << %ld" (reg d) (reg r1) (reg r2) (camlint_of_coqint n) - | Olea addr, args -> sprintf "%s=addr" (reg d) - | Omakelong, [r1;r2] -> sprintf "%s=makelong(%s,%s)" (reg d) (reg r1) (reg r2) - | Olowlong, [r1] -> sprintf "%s=lowlong(%s)" (reg d) (reg r1) - | Ohighlong, [r1] -> sprintf "%s=highlong(%s)" (reg d) (reg r1) - | Ocast32signed, [r1] -> sprintf "%s=long32signed(%s)" (reg d) (reg r1) - | Ocast32unsigned, [r1] -> sprintf "%s=long32unsigned(%s)" (reg d) (reg r1) - | Onegl, [r1] -> sprintf "%s=-l %s" (reg d) (reg r1) - | Osubl, [r1;r2] -> sprintf "%s=%s -l %s" (reg d) (reg r1) (reg r2) - | Omull, [r1;r2] -> sprintf "%s=%s *l %s" (reg d) (reg r1) (reg r2) - | Omullimm n, [r1] -> sprintf "%s=%s *l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Omullhs, [r1;r2] -> sprintf "%s=mullhs(%s,%s)" (reg d) (reg r1) (reg r2) - | Omullhu, [r1;r2] -> sprintf "%s=mullhu(%s,%s)" (reg d) (reg r1) (reg r2) - | Odivl, [r1;r2] -> sprintf "%s=%s /ls %s" (reg d) (reg r1) (reg r2) - | Odivlu, [r1;r2] -> sprintf "%s=%s /lu %s" (reg d) (reg r1) (reg r2) - | Omodl, [r1;r2] -> sprintf "%s=%s %%ls %s" (reg d) (reg r1) (reg r2) - | Omodlu, [r1;r2] -> sprintf "%s=%s %%lu %s" (reg d) (reg r1) (reg r2) - | Oandl, [r1;r2] -> sprintf "%s=%s &l %s" (reg d) (reg r1) (reg r2) - | Oandlimm n, [r1] -> sprintf "%s=%s &l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Oorl, [r1;r2] -> sprintf "%s=%s |l %s" (reg d) (reg r1) (reg r2) - | Oorlimm n, [r1] -> sprintf "%s=%s |l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Oxorl, [r1;r2] -> sprintf "%s=%s ^l %s" (reg d) (reg r1) (reg r2) - | Oxorlimm n, [r1] -> sprintf "%s=%s ^l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Onotl, [r1] -> sprintf "%s=notl(%s)" (reg d) (reg r1) - | Oshll, [r1;r2] -> sprintf "%s=%s < sprintf "%s=%s < sprintf "%s=%s >>ls %s" (reg d) (reg r1) (reg r2) - | Oshrlimm n, [r1] -> sprintf "%s=%s >>ls %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrxlimm n, [r1] -> sprintf "%s=%s >>lx %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrlu, [r1;r2] -> sprintf "%s=%s >>lu %s" (reg d) (reg r1) (reg r2) - | Oshrluimm n, [r1] -> sprintf "%s=%s >>lu %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Ororlimm n, [r1] -> sprintf "%s=%s rorl %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oleal addr, args -> sprintf "%s=addr" (reg d) - | Onegf, [r1] -> sprintf "%s=negf(%s)" (reg d) (reg r1) - | Oabsf, [r1] -> sprintf "%s=absf(%s)" (reg d) (reg r1) - | Oaddf, [r1;r2] -> sprintf "%s=%s +f %s" (reg d) (reg r1) (reg r2) - | Osubf, [r1;r2] -> sprintf "%s=%s -f %s" (reg d) (reg r1) (reg r2) - | Omulf, [r1;r2] -> sprintf "%s=%s *f %s" (reg d) (reg r1) (reg r2) - | Odivf, [r1;r2] -> sprintf "%s=%s /f %s" (reg d) (reg r1) (reg r2) - | Onegfs, [r1] -> sprintf "%s=negfs(%s)" (reg d) (reg r1) - | Oabsfs, [r1] -> sprintf "%s=absfs(%s)" (reg d) (reg r1) - | Oaddfs, [r1;r2] -> sprintf "%s=%s +fs %s" (reg d) (reg r1) (reg r2) - | Osubfs, [r1;r2] -> sprintf "%s=%s -fs %s" (reg d) (reg r1) (reg r2) - | Omulfs, [r1;r2] -> sprintf "%s=%s *fs %s" (reg d) (reg r1) (reg r2) - | Odivfs, [r1;r2] -> sprintf "%s=%s /fs %s" (reg d) (reg r1) (reg r2) - | Osingleoffloat, [r1] -> sprintf "%s=singleoffloat(%s)" (reg d) (reg r1) - | Ofloatofsingle, [r1] -> sprintf "%s=floatofsingle(%s)" (reg d) (reg r1) - | Ointoffloat, [r1] -> sprintf "%s=intoffloat(%s)" (reg d) (reg r1) - | Ofloatofint, [r1] -> sprintf "%s=floatofint(%s)" (reg d) (reg r1) - | Ointofsingle, [r1] -> sprintf "%s=intofsingle(%s)" (reg d) (reg r1) - | Osingleofint, [r1] -> sprintf "%s=singleofint(%s)" (reg d) (reg r1) - | Olongoffloat, [r1] -> sprintf "%s=longoffloat(%s)" (reg d) (reg r1) - | Ofloatoflong, [r1] -> sprintf "%s=floatoflong(%s)" (reg d) (reg r1) - | Olongofsingle, [r1] -> sprintf "%s=longofsingle(%s)" (reg d) (reg r1) - | Osingleoflong, [r1] -> sprintf "%s=singleoflong(%s)" (reg d) (reg r1) - | Ocmp c, args -> sprintf "%s=cmp" (reg d) - | Osel (c, ty), r1::r2::args -> sprintf "%s=sel" (reg d) - | _, _ -> sprintf "N/a") - -module DFGDot = Graph.Graphviz.Dot(struct - let graph_attributes _ = [] - let default_vertex_attributes _ = [] - let vertex_name = function (i, instr) -> sprintf "\"%d:%s\"" i (print_instr instr) - let vertex_attributes _ = [] - let get_subgraph _ = None - let default_edge_attributes _ = [] - let edge_attributes _ = [] - - include DFG - end) - -module DFGDfs = Graph.Traverse.Dfs(DFG) - -module IMap = Map.Make (struct - type t = int - - let compare = compare -end) - -let gen_vertex instrs i = (i, List.nth instrs i) - -(** The DFG type defines a list of instructions with their data dependencies as [edges], which are - the pairs of integers that represent the index of the instruction in the [nodes]. The edges - always point from left to right. *) - -let print_list f out_chan a = - fprintf out_chan "[ "; - List.iter (fprintf out_chan "%a " f) a; - fprintf out_chan "]" - -let print_tuple out_chan a = - let l, r = a in - fprintf out_chan "(%d,%d)" l r - -(*let print_dfg out_chan dfg = - fprintf out_chan "{ nodes = %a, edges = %a }" - (print_list PrintRTLBlockInstr.print_bblock_body) - dfg.nodes (print_list print_tuple) dfg.edges*) - -let print_dfg = DFGDot.output_graph - -let read_process command = - let buffer_size = 2048 in - let buffer = Buffer.create buffer_size in - let string = Bytes.create buffer_size in - let in_channel = Unix.open_process_in command in - let chars_read = ref 1 in - while !chars_read <> 0 do - chars_read := input in_channel string 0 buffer_size; - Buffer.add_substring buffer (Bytes.to_string string) 0 !chars_read - done; - ignore (Unix.close_process_in in_channel); - Buffer.contents buffer - -let comb_delay = function - | RBnop -> 0 - | RBop (_, op, _, _) -> - (match op with - | Omove -> 0 - | Ointconst _ -> 0 - | Olongconst _ -> 0 - | Ocast8signed -> 0 - | Ocast8unsigned -> 0 - | Ocast16signed -> 0 - | Ocast16unsigned -> 0 - | Oneg -> 0 - | Onot -> 0 - | Oor -> 0 - | Oorimm _ -> 0 - | Oand -> 0 - | Oandimm _ -> 0 - | Oxor -> 0 - | Oxorimm _ -> 0 - | Omul -> 8 - | Omulimm _ -> 8 - | Omulhs -> 8 - | Omulhu -> 8 - | Odiv -> 72 - | Odivu -> 72 - | Omod -> 72 - | Omodu -> 72 - | _ -> 1) - | _ -> 1 - -let pipeline_stages = function - | RBload _ -> 2 - | RBop (_, op, _, _) -> - (match op with - | Odiv -> 32 - | Odivu -> 32 - | Omod -> 32 - | Omodu -> 32 - | _ -> 0) - | _ -> 0 - -(** Add a dependency if it uses a register that was written to previously. *) -let add_dep map i tree dfg curr = - match PTree.get curr tree with - | None -> dfg - | Some ip -> - let ipv = (List.nth map ip) in - 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. - - This function only gathers the RAW constraints, and will therefore only be active for operations - that modify registers, which is this case only affects loads and operations. *) -let accumulate_RAW_deps map dfg curr = - let i, dst_map, graph = dfg in - let acc_dep_instruction rs dst = - ( i + 1, - PTree.set dst i dst_map, - List.fold_left (add_dep map i dst_map) graph rs - ) - in - let acc_dep_instruction_nodst rs = - ( i + 1, - dst_map, - List.fold_left (add_dep map i dst_map) graph rs) - in - 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) - -(** Finds the next write to the [dst] register. This is a small optimisation so that only one - dependency is generated for a data dependency. *) -let rec find_next_dst_write i dst i' curr = - let check_dst dst' curr' = - if dst = dst' then Some (i, i') - else find_next_dst_write i dst (i' + 1) curr' - in - match curr with - | [] -> None - | RBop (_, _, _, dst') :: curr' -> check_dst dst' curr' - | RBload (_, _, _, _, dst') :: curr' -> check_dst dst' curr' - | _ :: curr' -> find_next_dst_write i dst (i' + 1) curr' - -let rec find_all_next_dst_read i dst i' curr = - let check_dst rs curr' = - if List.exists (fun x -> x = dst) rs - then (i, i') :: find_all_next_dst_read i dst (i' + 1) curr' - else find_all_next_dst_read i dst (i' + 1) curr' - in - match curr with - | [] -> [] - | RBop (_, _, rs, _) :: curr' -> check_dst rs curr' - | RBload (_, _, _, rs, _) :: curr' -> check_dst rs curr' - | RBstore (_, _, _, rs, src) :: curr' -> check_dst (src :: rs) curr' - | RBnop :: curr' -> find_all_next_dst_read i dst (i' + 1) curr' - | RBsetpred (_, _, rs, _) :: curr' -> check_dst rs curr' - -let drop i lst = - let rec drop' i' lst' = - match lst' with - | _ :: ls -> if i' = i then ls else drop' (i' + 1) ls - | [] -> [] - in - if i = 0 then lst else drop' 1 lst - -let take i lst = - let rec take' i' lst' = - match lst' with - | l :: ls -> if i' = i then [ l ] else l :: take' (i' + 1) ls - | [] -> [] - in - if i = 0 then [] else take' 1 lst - -let rec next_store i = function - | [] -> None - | RBstore (_, _, _, _, _) :: _ -> Some i - | _ :: rst -> next_store (i + 1) rst - -let rec next_load i = function - | [] -> None - | RBload (_, _, _, _, _) :: _ -> Some i - | _ :: rst -> next_load (i + 1) rst - -let accumulate_RAW_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBload (_, _, _, _, _) -> ( - 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) ) - | _ -> dfg - -let accumulate_WAR_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBstore (_, _, _, _, _) -> ( - match next_load 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) ) - | _ -> dfg - -let accumulate_WAW_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | 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)) - | _ -> dfg - -(** Predicate dependencies. *) - -let rec in_predicate p p' = - match p' with - | Plit p'' -> P.to_int p = P.to_int (snd p'') - | Pand (p1, p2) -> in_predicate p p1 || in_predicate p p2 - | Por (p1, p2) -> in_predicate p p1 || in_predicate p p2 - | Ptrue -> false - | Pfalse -> false - -let get_predicate = function - | RBop (p, _, _, _) -> p - | RBload (p, _, _, _, _) -> p - | RBstore (p, _, _, _, _) -> p - | RBsetpred (p, _, _, _) -> p - | _ -> None - -let rec next_setpred p i = function - | [] -> None - | RBsetpred (_, _, _, p') :: rst -> - if in_predicate p' p then - Some i - else - next_setpred p (i + 1) rst - | _ :: rst -> next_setpred p (i + 1) rst - -let rec next_preduse p i instr= - let next p' rst = - if in_predicate p p' then - Some i - else - next_preduse p (i + 1) rst - in - match instr with - | [] -> None - | 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 = - let i, curr = curri in - match get_predicate curr with - | Some p -> ( - match next_setpred p 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) ) - | _ -> dfg - -let accumulate_WAR_pred_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBsetpred (_, _, _, p) -> ( - match next_preduse p 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) ) - | _ -> dfg - -let accumulate_WAW_pred_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBsetpred (_, _, _, p) -> ( - match next_setpred (Plit (true, p)) 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) ) - | _ -> dfg - -(** This function calculates the WAW dependencies, which happen when two writes are ordered one - after another and therefore have to be kept in that order. This accumulation might be redundant - if register renaming is done before hand, because then these dependencies can be avoided. *) -let accumulate_WAW_deps instrs dfg curri = - let i, curr = curri in - let dst_dep dst = - match find_next_dst_write i dst (i + 1) (drop (i + 1) instrs) with - | Some (a, b) -> DFG.add_edge dfg (gen_vertex instrs a) (gen_vertex instrs b) - | _ -> dfg - in - match curr with - | RBop (_, _, _, dst) -> dst_dep dst - | RBload (_, _, _, _, dst) -> dst_dep dst - | RBstore (_, _, _, _, _) -> ( - match next_store (i + 1) (drop (i + 1) instrs) with - | None -> dfg - | Some i' -> DFG.add_edge dfg (gen_vertex instrs i) (gen_vertex instrs i') ) - | _ -> dfg - -let accumulate_WAR_deps instrs dfg curri = - let i, curr = curri in - let dst_dep dst = - let dep_list = find_all_next_dst_read i dst 0 (take i instrs |> List.rev) - |> List.map (function (d, d') -> (i - d' - 1, d)) - in - List.fold_left (fun g -> - function (d, d') -> DFG.add_edge g (gen_vertex instrs d) (gen_vertex instrs d')) dfg dep_list - in - match curr with - | RBop (_, _, _, dst) -> dst_dep dst - | RBload (_, _, _, _, dst) -> dst_dep dst - | _ -> dfg - -let assigned_vars vars = function - | RBnop -> vars - | RBop (_, _, _, dst) -> dst :: vars - | RBload (_, _, _, _, dst) -> dst :: vars - | RBstore (_, _, _, _, _) -> vars - | RBsetpred (_, _, _, _) -> vars - -let get_pred = function - | RBnop -> None - | RBop (op, _, _, _) -> op - | RBload (op, _, _, _, _) -> op - | RBstore (op, _, _, _, _) -> op - | RBsetpred (_, _, _, _) -> None - -let independant_pred p p' = - match sat_pred_simple (Pand (p, p')) with - | None -> true - | _ -> false - -let check_dependent op1 op2 = - match op1, op2 with - | Some p, Some p' -> not (independant_pred p p') - | _, _ -> true - -let remove_unnecessary_deps graph = - let is_dependent v1 v2 g = - let (_, instr1) = v1 in - let (_, instr2) = v2 in - if check_dependent (get_pred instr1) (get_pred instr2) - then g - else DFG.remove_edge g v1 v2 - in - DFG.fold_edges is_dependent graph graph - -(** All the nodes in the DFG have to come after the source of the basic block, and should terminate - before the sink of the basic block. After that, there should be constraints for data - dependencies between nodes. *) -let gather_bb_constraints debug bb = - let ibody = List.mapi (fun i a -> (i, a)) bb.bb_body in - let dfg = List.fold_left (fun dfg v -> DFG.add_vertex dfg v) DFG.empty ibody in - let _, _, dfg' = - List.fold_left (accumulate_RAW_deps ibody) - (0, PTree.empty, dfg) - bb.bb_body - in - let dfg'' = List.fold_left (fun dfg f -> List.fold_left (f bb.bb_body) dfg ibody) dfg' - [ accumulate_WAW_deps; - accumulate_WAR_deps; - accumulate_RAW_mem_deps; - accumulate_WAR_mem_deps; - accumulate_WAW_mem_deps; - accumulate_RAW_pred_deps; - accumulate_WAR_pred_deps; - accumulate_WAW_pred_deps - ] - in - let dfg''' = remove_unnecessary_deps dfg'' in - (List.length bb.bb_body, dfg''', successors_instr bb.bb_exit) - -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 - then G.add_edge_e g (encode_bb n 0, 0, encode_var n (fst v1) 0) - else g) |> - (fun g' -> - if DFG.out_degree dfg v1 = 0 - 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, 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 = - List.fold_left (fun g n' -> - G.add_edge_e g (encode_bb n 1, -1, encode_bb n' 0) - ) constr succs - -module BFDFG = Graph.Path.BellmanFord(DFG)(struct - include DFG - type t = int - let weight = DFG.E.label - let compare = compare - let add = (+) - let zero = 0 - end) - -module TopoDFG = Graph.Topological.Make(DFG) - -let negate_graph constr = - DFG.fold_edges_e (function (v1, e, v2) -> fun g -> - DFG.add_edge_e g (v1, -e, v2) - ) constr DFG.empty - -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 - |> 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 maxf))) - 1), - encode_var n (fst v') 0) - ) constr (DFG.fold_vertex (fun v l -> - 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 = - | Mem - | SDiv - | UDiv - -type resources = { - res_mem: DFG.V.t list; - res_udiv: DFG.V.t list; - res_sdiv: DFG.V.t list; -} - -let find_resource = function - | RBload _ -> Some Mem - | RBstore _ -> Some Mem - | RBop (_, op, _, _) -> - ( match op with - | Odiv -> Some SDiv - | Odivu -> Some UDiv - | Omod -> Some SDiv - | Omodu -> Some UDiv - | _ -> None ) - | _ -> None - -let add_resource_constr n dfg constr = - let res = TopoDFG.fold (function (i, instr) -> - function {res_mem = ml; res_sdiv = sdl; res_udiv = udl} as r -> - match find_resource instr with - | Some SDiv -> {r with res_sdiv = (i, instr) :: sdl} - | Some UDiv -> {r with res_udiv = (i, instr) :: udl} - | Some Mem -> {r with res_mem = (i, instr) :: ml} - | None -> r - ) dfg {res_mem = []; res_sdiv = []; res_udiv = []} - in - let get_constraints l g = List.fold_left (fun gv v' -> - match gv with - | (g, None) -> (g, Some v') - | (g, Some v) -> - (G.add_edge_e g (encode_var n (fst v) 0, -1, encode_var n (fst v') 0), Some v') - ) (g, None) l |> fst - in - get_constraints (List.rev res.res_mem) constr - |> get_constraints (List.rev res.res_udiv) - |> get_constraints (List.rev res.res_sdiv) - -let gather_cfg_constraints c constr curr = - let (n, dfg) = curr in - match PTree.get (P.of_int n) c with - | 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 - |> List.filter (fun n' -> n' < n)) - |> add_cycle_constr 8 n dfg - |> add_resource_constr n dfg - -let rec intersperse s = function - | [] -> [] - | [ a ] -> [ a ] - | x :: xs -> x :: s :: intersperse s xs - -let print_objective constr = - let vars = G.fold_vertex (fun v1 l -> - match v1 with - | { sv_type = VarType _; sv_num = 0 } -> print_sv v1 :: l - | _ -> l - ) constr [] - in - "min: " ^ String.concat "" (intersperse " + " vars) ^ ";\n" - -let print_lp constr = - print_objective constr ^ - (G.fold_edges_e (function (e1, w, e2) -> fun s -> - s ^ sprintf "%s - %s <= %d;\n" (print_sv e1) (print_sv e2) w - ) constr "" |> - G.fold_vertex (fun v1 s -> - s ^ sprintf "int %s;\n" (print_sv v1) - ) constr) - -let update_schedule v = function Some l -> Some (v :: l) | None -> Some [ v ] - -let parse_soln (tree, bbtree) s = - let r = Str.regexp "var\\([0-9]+\\)n\\([0-9]+\\)_0[ ]+\\([0-9]+\\)" in - 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 (fn, oc) = Filename.open_temp_file "vericert_" "_lp_solve" in - fprintf oc "%s\n" (print_lp constr); - close_out oc; - - let res = Str.split (Str.regexp_string "\n") (read_process ("lp_solve " ^ fn)) - |> drop 3 - |> List.fold_left parse_soln (IMap.empty, 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 - List.fold_left (fun g v -> - List.fold_left (fun g' v' -> - let edges = DFG.find_all_edges dfg v v' in - List.fold_left DFG.add_edge_e g' edges - ) g l - ) dfg' l - -let rec all_successors dfg v = - List.concat (List.fold_left (fun l v -> - all_successors dfg v :: l - ) [] (DFG.succ dfg v)) - -let order_instr dfg = - DFG.fold_vertex (fun v li -> - if DFG.in_degree dfg v = 0 - then (List.map snd (v :: all_successors dfg v)) :: li - else li - ) dfg [] - -let combine_bb_schedule schedule s = - let i, st = s in - IMap.update st (update_schedule i) schedule - -(**let add_el dfg i l = - List.*) - -let check_in el = - List.exists (List.exists ((=) el)) - -let all_dfs dfg = - let roots = DFG.fold_vertex (fun v li -> - if DFG.in_degree dfg v = 0 then v :: li else li - ) dfg [] in - let dfg' = DFG.fold_edges (fun v1 v2 g -> DFG.add_edge g v2 v1) dfg dfg in - List.fold_left (fun a el -> - if check_in el a then a else - (DFGDfs.fold_component (fun v l -> v :: l) [] dfg' el) :: a) [] roots - -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 = - 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 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 - |> 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 -> - all_dfs x - |> List.map (subgraph x) - |> List.map (fun y -> - TopoDFG.fold (fun i l -> snd i :: l) y [] - |> List.rev))) body2 - (*|> (fun x -> TopoDFG.fold (fun i l -> snd i :: l) x []) - |> List.rev) body2*) - in - { bb_body = final_body2; - bb_exit = ctrl_flow - } - in - PTree.map f c - -let schedule entry (c : RTLBlock.bb RTLBlockInstr.code) = - let debug = true in - let transf_graph (_, dfg, _) = dfg in - let c' = PTree.map1 (fun x -> gather_bb_constraints false x |> transf_graph) c in - (*let _ = if debug then PTree.map (fun r o -> printf "##### %d #####\n%a\n\n" (P.to_int r) print_dfg o) c' else PTree.empty in*) - let cgraph = PTree.elements c' - |> List.map (function (x, y) -> (P.to_int x, y)) - |> List.fold_left (gather_cfg_constraints c) G.empty - in - let graph = open_out "constr_graph.dot" in - fprintf graph "%a\n" GDot.output_graph cgraph; - close_out graph; - let schedule' = solve_constraints cgraph in - (**IMap.iter (fun a b -> printf "##### %d #####\n%a\n\n" a (print_list print_tuple) b) schedule';*) - (**printf "Schedule: %a\n" (fun a x -> IMap.iter (fun d -> fprintf a "%d: %a\n" d (print_list print_tuple)) x) schedule';*) - transf_rtlpar c c' schedule' - -let rec find_reachable_states c e = - match PTree.get e c with - | Some { bb_exit = ex; _ } -> - e :: List.fold_left (fun x a -> List.concat [x; find_reachable_states c a]) [] - (successors_instr ex |> List.filter (fun x -> P.lt x e)) - | None -> assert false - -let add_to_tree c nt i = - match PTree.get i c with - | Some p -> PTree.set i p nt - | None -> assert false - -let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = - let scheduled = schedule f.fn_entrypoint f.fn_code in - let reachable = find_reachable_states scheduled f.fn_entrypoint - |> List.to_seq |> SS.of_seq |> SS.to_seq |> List.of_seq in - { fn_sig = f.fn_sig; - fn_params = f.fn_params; - fn_stacksize = f.fn_stacksize; - fn_code = scheduled (*List.fold_left (add_to_tree scheduled) PTree.empty reachable*); - fn_entrypoint = f.fn_entrypoint - } -#+end_src - -* RTLBlockgen -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgen.v -:END: - -Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]]. - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblockgen-imports -#+begin_src coq -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. -#+end_src - -#+name: rtlblockgen-equalities-insert -#+begin_src coq -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. -#+end_src - -#+name: rtlblockgen-main -#+begin_src coq -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 - | 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. -#+end_src - -* RTLBlockgenproof -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -** Imports - -#+name: rtlblockgenproof-imports -#+begin_src coq -Require compcert.backend.RTL. -Require Import compcert.common.AST. -Require Import compcert.lib.Maps. - -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLBlockgen. -#+end_src - -** Match states - -The ~match_states~ predicate describes which states are equivalent between the two languages, in this -case ~RTL~ and ~RTLBlock~. - -#+name: rtlblockgenproof-match-states -#+begin_src coq -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). -#+end_src - -** Correctness - -#+name: rtlblockgenproof-correctness -#+begin_src coq -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. -#+end_src -* RTLPargen -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargen.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpargen-main -#+begin_src coq -Require Import compcert.backend.Registers. -Require Import compcert.common.AST. -Require Import compcert.common.Globalenvs. -Require Import compcert.common.Memory. -Require Import compcert.common.Values. -Require Import compcert.lib.Floats. -Require Import compcert.lib.Integers. -Require Import compcert.lib.Maps. -Require compcert.verilog.Op. - -Require Import vericert.common.Vericertlib. -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLPar. -Require Import vericert.hls.RTLBlockInstr. -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. - -(** ** 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) - end. - -Fixpoint replicate {A} (n: nat) (l: A) := - match n with - | O => nil - | S n => l :: replicate n l - end. - -Definition merge''' x y := - match x, y with - | Some p1, Some p2 => Some (Pand p1 p2) - | Some p, None | None, Some p => Some p - | None, None => None - end. - -Definition merge'' x := - match x with - | ((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 := - match pa, pf with - | (p, a), (p', f) => (merge''' p p', f a) - end. - -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. - -(*map (fun x => (fst x, Econs (snd x) Enil)) pel*) -Definition merge' (pel: pred_expr) (tpel: predicated expression_list) := - predicated_map (uncurry Econs) (predicated_prod pel tpel). - -Fixpoint merge (pel: list pred_expr): predicated expression_list := - match pel with - | nil => NE.singleton (T, Enil) - | a :: b => merge' a (merge b) - end. - -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 := - 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 := - 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 := - NE.map (fun x => (fst x, (snd x) pa pb pc)) pf. - -Definition predicated_from_opt {A: Type} (p: option pred_op) (a: A) := - match p with - | Some p' => NE.singleton (p', a) - | None => NE.singleton (T, a) - end. - -#[local] Open Scope non_empty_scope. -#[local] Open Scope pred_op. - -Fixpoint NEfold_left {A B} (f: A -> B -> A) (l: NE.non_empty B) (a: A) : A := - match l with - | NE.singleton a' => f a a' - | a' ::| b => NEfold_left f b (f a a') - end. - -Fixpoint NEapp {A} (l m: NE.non_empty A) := - match l with - | NE.singleton a => a ::| m - | a ::| b => a ::| NEapp b m - end. - -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. - -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 - -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 - | RBop p op rl r => - f # (Reg r) <- - (app_predicated p - (f # (Reg r)) - (map_predicated (pred_ret (Eop op)) (merge (list_translation rl f)))) - | RBload p chunk addr rl r => - f # (Reg r) <- - (app_predicated p - (f # (Reg r)) - (map_predicated - (map_predicated (pred_ret (Eload chunk addr)) (merge (list_translation rl f))) - (f # Mem))) - | RBstore p chunk addr rl r => - f # Mem <- - (app_predicated p - (f # Mem) - (map_predicated - (map_predicated - (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)))) - 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. - -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. - -(** 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. *) - -Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := - match bb with - | nil => - match bbt with - | nil => true - | _ => false - end - | _ => true - end. - -Definition schedule_oracle (bb: RTLBlock.bblock) (bbt: RTLPar.bblock) : bool := - check (abstract_sequence empty (bb_body bb)) - (abstract_sequence empty (concat (concat (bb_body bbt)))) && - check_control_flow_instr (bb_exit bb) (bb_exit bbt) && - empty_trees (bb_body bb) (bb_body bbt). - -Definition check_scheduled_trees := beq2 schedule_oracle. - -Ltac solve_scheduled_trees_correct := - intros; unfold check_scheduled_trees in *; - match goal with - | [ H: context[beq2 _ _ _], x: positive |- _ ] => - rewrite beq2_correct in H; specialize (H x) - end; repeat destruct_match; crush. - -Lemma check_scheduled_trees_correct: - forall f1 f2 x y1, - check_scheduled_trees f1 f2 = true -> - PTree.get x f1 = Some y1 -> - exists y2, PTree.get x f2 = Some y2 /\ schedule_oracle y1 y2 = true. -Proof. solve_scheduled_trees_correct; eexists; crush. Qed. - -Lemma check_scheduled_trees_correct2: - forall f1 f2 x, - check_scheduled_trees f1 f2 = true -> - PTree.get x f1 = None -> - PTree.get x f2 = None. -Proof. solve_scheduled_trees_correct. Qed. - -(** ** Top-level Functions *) - -Parameter schedule : RTLBlock.function -> 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) - f.(fn_params) - f.(fn_stacksize) - tfcode - f.(fn_entrypoint)) - else - 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. -#+end_src -* RTLPargenproof -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargenproof.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpargenproof-imports -#+begin_src coq -Require Import compcert.backend.Registers. -Require Import compcert.common.AST. -Require Import compcert.common.Errors. -Require Import compcert.common.Linking. -Require Import compcert.common.Globalenvs. -Require Import compcert.common.Memory. -Require Import compcert.common.Values. -Require Import compcert.lib.Maps. - -Require Import vericert.common.Vericertlib. -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLPar. -Require Import vericert.hls.RTLBlockInstr. -Require Import vericert.hls.RTLPargen. -Require Import vericert.hls.Predicate. -Require Import vericert.hls.Abstr. - -#[local] Open Scope positive. -#[local] Open Scope forest. -#[local] Open Scope pred_op. -#+end_src - -** RTLBlock to abstract translation - -Correctness of translation from RTLBlock to the abstract interpretation language. - -#+name: rtlpargenproof-main -#+begin_src coq -(*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. - -Inductive state_lessdef : instr_state -> instr_state -> Prop := - state_lessdef_intro : - forall rs1 rs2 m1, - (forall x, rs1 !! x = rs2 !! x) -> - state_lessdef (mk_instr_state rs1 m1) (mk_instr_state rs2 m1). - -Ltac inv_simp := - repeat match goal with - | H: exists _, _ |- _ => inv H - end; simplify. - -*) - -Definition check_dest i r' := - match i with - | RBop p op rl r => (r =? r')%positive - | RBload p chunk addr rl r => (r =? r')%positive - | _ => false - end. - -Lemma check_dest_dec i r : {check_dest i r = true} + {check_dest i r = false}. -Proof. destruct (check_dest i r); tauto. Qed. - -Fixpoint check_dest_l l r := - match l with - | nil => false - | a :: b => check_dest a r || check_dest_l b r - end. - -Lemma check_dest_l_forall : - forall l r, - check_dest_l l r = false -> - Forall (fun x => check_dest x r = false) l. -Proof. induction l; crush. Qed. - -Lemma check_dest_l_dec i r : {check_dest_l i r = true} + {check_dest_l i r = false}. -Proof. destruct (check_dest_l i r); tauto. Qed. - -Lemma check_dest_update : - forall f i r, - check_dest i r = false -> - (update f i) # (Reg r) = f # (Reg r). -Proof. - destruct i; crush; try apply Pos.eqb_neq in H; apply genmap1; crush. -Qed. - -Lemma check_dest_l_forall2 : - forall l r, - Forall (fun x => check_dest x r = false) l -> - check_dest_l l r = false. -Proof. - induction l; crush. - inv H. apply orb_false_intro; crush. -Qed. - -Lemma check_dest_l_ex2 : - forall l r, - (exists a, In a l /\ check_dest a r = true) -> - check_dest_l l r = true. -Proof. - induction l; crush. - specialize (IHl r). inv H. - apply orb_true_intro; crush. - apply orb_true_intro; crush. - right. apply IHl. exists x. auto. -Qed. - -Lemma check_list_l_false : - forall l x r, - check_dest_l (l ++ x :: nil) r = false -> - check_dest_l l r = false /\ check_dest x r = false. -Proof. - simplify. - apply check_dest_l_forall in H. apply Forall_app in H. - simplify. apply check_dest_l_forall2; auto. - apply check_dest_l_forall in H. apply Forall_app in H. - simplify. inv H1. auto. -Qed. - -Lemma check_dest_l_ex : - forall l r, - check_dest_l l r = true -> - exists a, In a l /\ check_dest a r = true. -Proof. - induction l; crush. - destruct (check_dest a r) eqn:?; try solve [econstructor; crush]. - simplify. - exploit IHl. apply H. simplify. econstructor. simplify. right. eassumption. - auto. -Qed. - -Lemma check_list_l_true : - forall l x r, - check_dest_l (l ++ x :: nil) r = true -> - check_dest_l l r = true \/ check_dest x r = true. -Proof. - simplify. - apply check_dest_l_ex in H; simplify. - apply in_app_or in H. inv H. left. - apply check_dest_l_ex2. exists x0. auto. - inv H0; auto. -Qed. - -Lemma check_dest_l_dec2 l r : - {Forall (fun x => check_dest x r = false) l} - + {exists a, In a l /\ check_dest a r = true}. -Proof. - destruct (check_dest_l_dec l r); [right | left]; - auto using check_dest_l_ex, check_dest_l_forall. -Qed. - -Lemma abstr_comp : - forall l i f x x0, - abstract_sequence f (l ++ i :: nil) = x -> - abstract_sequence f l = x0 -> - x = update x0 i. -Proof. induction l; intros; crush; eapply IHl; eauto. Qed. - -(* - -Lemma gen_list_base: - forall FF ge sp l rs exps st1, - (forall x, @sem_value FF ge sp st1 (exps # (Reg x)) (rs !! x)) -> - sem_val_list ge sp st1 (list_translation l exps) rs ## l. -Proof. - induction l. - intros. simpl. constructor. - intros. simpl. eapply Scons; eauto. -Qed. - -Lemma check_dest_update2 : - forall f r rl op p, - (update f (RBop p op rl r)) # (Reg r) = Eop op (list_translation rl f) (f # Mem). -Proof. crush; rewrite map2; auto. Qed. - -Lemma check_dest_update3 : - forall f r rl p addr chunk, - (update f (RBload p chunk addr rl r)) # (Reg r) = Eload chunk addr (list_translation rl f) (f # Mem). -Proof. crush; rewrite map2; auto. Qed. - -Lemma abstract_seq_correct_aux: - forall FF ge sp i st1 st2 st3 f, - @step_instr FF ge sp st3 i st2 -> - sem ge sp st1 f st3 -> - sem ge sp st1 (update f i) st2. -Proof. - intros; inv H; simplify. - { simplify; eauto. } (*apply match_states_refl. }*) - { inv H0. inv H6. destruct st1. econstructor. simplify. - constructor. intros. - destruct (resource_eq (Reg res) (Reg x)). inv e. - rewrite map2. econstructor. eassumption. apply gen_list_base; eauto. - rewrite Regmap.gss. eauto. - assert (res <> x). { unfold not in *. intros. apply n. rewrite H0. auto. } - rewrite Regmap.gso by auto. - rewrite genmap1 by auto. auto. - - rewrite genmap1; crush. } - { inv H0. inv H7. constructor. constructor. intros. - destruct (Pos.eq_dec dst x); subst. - rewrite map2. econstructor; eauto. - apply gen_list_base. auto. rewrite Regmap.gss. auto. - rewrite genmap1. rewrite Regmap.gso by auto. auto. - unfold not in *; intros. inv H0. auto. - rewrite genmap1; crush. - } - { inv H0. inv H7. constructor. constructor; intros. - rewrite genmap1; crush. - rewrite map2. econstructor; eauto. - apply gen_list_base; auto. - } -Qed. - -Lemma regmap_list_equiv : - forall A (rs1: Regmap.t A) rs2, - (forall x, rs1 !! x = rs2 !! x) -> - forall rl, rs1##rl = rs2##rl. -Proof. induction rl; crush. Qed. - -Lemma sem_update_Op : - forall A ge sp st f st' r l o0 o m rs v, - @sem A ge sp st f st' -> - Op.eval_operation ge sp o0 rs ## l m = Some v -> - match_states st' (mk_instr_state rs m) -> - exists tst, - sem ge sp st (update f (RBop o o0 l r)) tst /\ match_states (mk_instr_state (Regmap.set r v rs) m) tst. -Proof. - intros. inv H1. simplify. - destruct st. - econstructor. simplify. - { constructor. - { constructor. intros. destruct (Pos.eq_dec x r); subst. - { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. - { inv H9. eapply gen_list_base; eauto. } - { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } - { inv H. rewrite genmap1; crush. eauto. } } - { constructor; eauto. intros. - destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed. - -Lemma sem_update_load : - forall A ge sp st f st' r o m a l m0 rs v a0, - @sem A ge sp st f st' -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.loadv m m0 a0 = Some v -> - match_states st' (mk_instr_state rs m0) -> - exists tst : instr_state, - sem ge sp st (update f (RBload o m a l r)) tst - /\ match_states (mk_instr_state (Regmap.set r v rs) m0) tst. -Proof. - intros. inv H2. pose proof H. inv H. inv H9. - destruct st. - econstructor; simplify. - { constructor. - { constructor. intros. - destruct (Pos.eq_dec x r); subst. - { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. - rewrite <- H6. eauto. - instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. auto. } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. } } - { rewrite genmap1; crush. eauto. } } - { constructor; auto; intros. destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed. - -Lemma sem_update_store : - forall A ge sp a0 m a l r o f st m' rs m0 st', - @sem A ge sp st f st' -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.storev m m0 a0 rs !! r = Some m' -> - match_states st' (mk_instr_state rs m0) -> - exists tst, sem ge sp st (update f (RBstore o m a l r)) tst - /\ match_states (mk_instr_state rs m') tst. -Proof. - intros. inv H2. pose proof H. inv H. inv H9. - destruct st. - econstructor; simplify. - { econstructor. - { econstructor; intros. rewrite genmap1; crush. } - { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. rewrite <- H6. - eauto. specialize (H6 r). rewrite H6. eauto. } } - { econstructor; eauto. } -Qed. - -Lemma sem_update : - forall A ge sp st x st' st'' st''' f, - sem ge sp st f st' -> - match_states st' st''' -> - @step_instr A ge sp st''' x st'' -> - exists tst, sem ge sp st (update f x) tst /\ match_states st'' tst. -Proof. - intros. destruct x; inv H1. - { econstructor. split. - apply sem_update_RBnop. eassumption. - apply match_states_commut. auto. } - { eapply sem_update_Op; eauto. } - { eapply sem_update_load; eauto. } - { eapply sem_update_store; eauto. } -Qed. - -Lemma sem_update2_Op : - forall A ge sp st f r l o0 o m rs v, - @sem A ge sp st f (mk_instr_state rs m) -> - Op.eval_operation ge sp o0 rs ## l m = Some v -> - sem ge sp st (update f (RBop o o0 l r)) (mk_instr_state (Regmap.set r v rs) m). -Proof. - intros. destruct st. constructor. - inv H. inv H6. - { constructor; intros. simplify. - destruct (Pos.eq_dec r x); subst. - { rewrite map2. econstructor. eauto. - apply gen_list_base. eauto. - rewrite Regmap.gss. auto. } - { rewrite genmap1; crush. rewrite Regmap.gso; auto. } } - { simplify. rewrite genmap1; crush. inv H. eauto. } -Qed. - -Lemma sem_update2_load : - forall A ge sp st f r o m a l m0 rs v a0, - @sem A ge sp st f (mk_instr_state rs m0) -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.loadv m m0 a0 = Some v -> - sem ge sp st (update f (RBload o m a l r)) (mk_instr_state (Regmap.set r v rs) m0). -Proof. - intros. simplify. inv H. inv H7. constructor. - { constructor; intros. destruct (Pos.eq_dec r x); subst. - { rewrite map2. rewrite Regmap.gss. econstructor; eauto. - apply gen_list_base; eauto. } - { rewrite genmap1; crush. rewrite Regmap.gso; eauto. } - } - { simplify. rewrite genmap1; crush. } -Qed. - -Lemma sem_update2_store : - forall A ge sp a0 m a l r o f st m' rs m0, - @sem A ge sp st f (mk_instr_state rs m0) -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.storev m m0 a0 rs !! r = Some m' -> - sem ge sp st (update f (RBstore o m a l r)) (mk_instr_state rs m'). -Proof. - intros. simplify. inv H. inv H7. constructor; simplify. - { econstructor; intros. rewrite genmap1; crush. } - { rewrite map2. econstructor; eauto. apply gen_list_base; eauto. } -Qed. - -Lemma sem_update2 : - forall A ge sp st x st' st'' f, - sem ge sp st f st' -> - @step_instr A ge sp st' x st'' -> - sem ge sp st (update f x) st''. -Proof. - intros. - destruct x; inv H0; - eauto using sem_update_RBnop, sem_update2_Op, sem_update2_load, sem_update2_store. -Qed. - -Lemma abstr_sem_val_mem : - forall A B ge tge st tst sp a, - ge_preserved ge tge -> - forall v m, - (@sem_mem A ge sp st a m /\ match_states st tst -> @sem_mem B tge sp tst a m) /\ - (@sem_value A ge sp st a v /\ match_states st tst -> @sem_value B tge sp tst a v). -Proof. - intros * H. - apply expression_ind2 with - - (P := fun (e1: expression) => - forall v m, - (@sem_mem A ge sp st e1 m /\ match_states st tst -> @sem_mem B tge sp tst e1 m) /\ - (@sem_value A ge sp st e1 v /\ match_states st tst -> @sem_value B tge sp tst e1 v)) - - (P0 := fun (e1: expression_list) => - forall lv, @sem_val_list A ge sp st e1 lv /\ match_states st tst -> @sem_val_list B tge sp tst e1 lv); - simplify; intros; simplify. - { inv H1. inv H2. constructor. } - { inv H2. inv H1. rewrite H0. constructor. } - { inv H3. } - { inv H3. inv H4. econstructor. apply H1; auto. simplify. eauto. constructor. auto. auto. - apply H0; simplify; eauto. constructor; eauto. - unfold ge_preserved in *. simplify. rewrite <- H2. auto. - } - { inv H3. } - { inv H3. inv H4. econstructor. apply H1; eauto; simplify; eauto. constructor; eauto. - apply H0; simplify; eauto. constructor; eauto. - inv H. rewrite <- H4. eauto. - auto. - } - { inv H4. inv H5. econstructor. apply H0; eauto. simplify; eauto. constructor; eauto. - apply H2; eauto. simplify; eauto. constructor; eauto. - apply H1; eauto. simplify; eauto. constructor; eauto. - inv H. rewrite <- H5. eauto. auto. - } - { inv H4. } - { inv H1. constructor. } - { inv H3. constructor; auto. apply H0; eauto. apply Mem.empty. } -Qed. - -Lemma abstr_sem_value : - forall a A B ge tge sp st tst v, - @sem_value A ge sp st a v -> - ge_preserved ge tge -> - match_states st tst -> - @sem_value B tge sp tst a v. -Proof. intros; eapply abstr_sem_val_mem; eauto; apply Mem.empty. Qed. - -Lemma abstr_sem_mem : - forall a A B ge tge sp st tst v, - @sem_mem A ge sp st a v -> - ge_preserved ge tge -> - match_states st tst -> - @sem_mem B tge sp tst a v. -Proof. intros; eapply abstr_sem_val_mem; eauto. Qed. - -Lemma abstr_sem_regset : - forall a a' A B ge tge sp st tst rs, - @sem_regset A ge sp st a rs -> - ge_preserved ge tge -> - (forall x, a # x = a' # x) -> - match_states st tst -> - exists rs', @sem_regset B tge sp tst a' rs' /\ (forall x, rs !! x = rs' !! x). -Proof. - inversion 1; intros. - inv H7. - econstructor. simplify. econstructor. intros. - eapply abstr_sem_value; eauto. rewrite <- H6. - eapply H0. constructor; eauto. - auto. -Qed. - -Lemma abstr_sem : - forall a a' A B ge tge sp st tst st', - @sem A ge sp st a st' -> - ge_preserved ge tge -> - (forall x, a # x = a' # x) -> - match_states st tst -> - exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. -Proof. - inversion 1; subst; intros. - inversion H4; subst. - exploit abstr_sem_regset; eauto; inv_simp. - do 3 econstructor; eauto. - rewrite <- H3. - eapply abstr_sem_mem; eauto. -Qed. - -Lemma abstract_execution_correct': - forall A B ge tge sp st' a a' st tst, - @sem A ge sp st a st' -> - ge_preserved ge tge -> - check a a' = true -> - match_states st tst -> - exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. -Proof. - intros; - pose proof (check_correct a a' H1); - eapply abstr_sem; eauto. -Qed. - -Lemma states_match : - forall st1 st2 st3 st4, - match_states st1 st2 -> - match_states st2 st3 -> - match_states st3 st4 -> - match_states st1 st4. -Proof. - intros * H1 H2 H3; destruct st1; destruct st2; destruct st3; destruct st4. - inv H1. inv H2. inv H3; constructor. - unfold regs_lessdef in *. intros. - repeat match goal with - | H: forall _, _, r : positive |- _ => specialize (H r) - end. - congruence. - auto. -Qed. - -Lemma step_instr_block_same : - forall ge sp st st', - step_instr_block ge sp st nil st' -> - st = st'. -Proof. inversion 1; auto. Qed. - -Lemma step_instr_seq_same : - forall ge sp st st', - step_instr_seq ge sp st nil st' -> - st = st'. -Proof. inversion 1; auto. Qed. - -Lemma sem_update' : - forall A ge sp st a x st', - sem ge sp st (update (abstract_sequence empty a) x) st' -> - exists st'', - @step_instr A ge sp st'' x st' /\ - sem ge sp st (abstract_sequence empty a) st''. -Proof. - Admitted. - -Lemma rtlpar_trans_correct : - forall bb ge sp sem_st' sem_st st, - sem ge sp sem_st (abstract_sequence empty (concat (concat bb))) sem_st' -> - match_states sem_st st -> - exists st', RTLPar.step_instr_block ge sp st bb st' - /\ match_states sem_st' st'. -Proof. - induction bb using rev_ind. - { repeat econstructor. eapply abstract_interp_empty3 in H. - inv H. inv H0. constructor; congruence. } - { simplify. inv H0. repeat rewrite concat_app in H. simplify. - rewrite app_nil_r in H. - exploit sem_separate; eauto; inv_simp. - repeat econstructor. admit. admit. - } -Admitted. - -(*Lemma abstract_execution_correct_ld: - forall bb bb' cfi ge tge sp st st' tst, - RTLBlock.step_instr_list ge sp st bb st' -> - ge_preserved ge tge -> - schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi) = true -> - match_states_ld st tst -> - exists tst', RTLPar.step_instr_block tge sp tst bb' tst' - /\ match_states st' tst'. -Proof. - intros.*) -*) - -Lemma match_states_list : - forall A (rs: Regmap.t A) rs', - (forall r, rs !! r = rs' !! r) -> - forall l, rs ## l = rs' ## l. -Proof. induction l; crush. Qed. - -Lemma PTree_matches : - forall A (v: A) res rs rs', - (forall r, rs !! r = rs' !! r) -> - forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x. -Proof. - intros; destruct (Pos.eq_dec x res); subst; - [ repeat rewrite Regmap.gss by auto - | repeat rewrite Regmap.gso by auto ]; auto. -Qed. - -Lemma abstract_interp_empty3 : - forall A ctx st', - @sem A ctx empty st' -> match_states (ctx_is ctx) st'. -Proof. - inversion 1; subst; simplify. destruct ctx. - destruct ctx_is. - constructor; intros. - - inv H0. specialize (H3 x). inv H3. inv H8. reflexivity. - - inv H1. specialize (H3 x). inv H3. inv H8. reflexivity. - - inv H2. inv H8. reflexivity. -Qed. - -Lemma step_instr_matches : - forall A a ge sp st st', - @step_instr A ge sp st a st' -> - forall tst, - match_states st tst -> - exists tst', step_instr ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction 1; simplify; - match goal with H: match_states _ _ |- _ => inv H end; - try solve [repeat econstructor; try erewrite match_states_list; - try apply PTree_matches; eauto; - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end]. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - apply PTree_matches; assumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - econstructor; auto. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - apply PTree_matches; assumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - econstructor; auto. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - erewrite <- eval_predf_pr_equiv; eassumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - erewrite <- eval_predf_pr_equiv; eassumption. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - - admit. Admitted. - -Lemma step_instr_list_matches : - forall a ge sp st st', - step_instr_list ge sp st a st' -> - forall tst, match_states st tst -> - exists tst', step_instr_list ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction a; intros; inv H; - try (exploit step_instr_matches; eauto; []; simplify; - exploit IHa; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma step_instr_seq_matches : - forall a ge sp st st', - step_instr_seq ge sp st a st' -> - forall tst, match_states st tst -> - exists tst', step_instr_seq ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction a; intros; inv H; - try (exploit step_instr_list_matches; eauto; []; simplify; - exploit IHa; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma step_instr_block_matches : - forall bb ge sp st st', - step_instr_block ge sp st bb st' -> - forall tst, match_states st tst -> - exists tst', step_instr_block ge sp tst bb tst' - /\ match_states st' tst'. -Proof. - induction bb; intros; inv H; - try (exploit step_instr_seq_matches; eauto; []; simplify; - exploit IHbb; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma rtlblock_trans_correct' : - forall bb ge sp st x st'', - RTLBlock.step_instr_list ge sp st (bb ++ x :: nil) st'' -> - exists st', RTLBlock.step_instr_list ge sp st bb st' - /\ step_instr ge sp st' x st''. -Proof. - induction bb. - crush. exists st. - split. constructor. inv H. inv H6. auto. - crush. inv H. exploit IHbb. eassumption. simplify. - econstructor. split. - econstructor; eauto. eauto. -Qed. - -Lemma abstract_interp_empty A st : @sem A st empty (ctx_is st). -Proof. destruct st, ctx_is. simpl. repeat econstructor. Qed. - -Lemma abstract_seq : - forall l f i, - abstract_sequence f (l ++ i :: nil) = update (abstract_sequence f l) i. -Proof. induction l; crush. Qed. - -Lemma abstract_sequence_update : - forall l r f, - check_dest_l l r = false -> - (abstract_sequence f l) # (Reg r) = f # (Reg r). -Proof. - induction l using rev_ind; crush. - rewrite abstract_seq. rewrite check_dest_update. apply IHl. - apply check_list_l_false in H. tauto. - apply check_list_l_false in H. tauto. -Qed. - -(*Lemma sem_separate : - forall A ctx b a st', - sem ctx (abstract_sequence empty (a ++ b)) st' -> - exists st'', - @sem A ctx (abstract_sequence empty a) st'' - /\ @sem A (mk_ctx st'' (ctx_sp ctx) (ctx_ge ctx)) (abstract_sequence empty b) st'. -Proof. - induction b using rev_ind; simplify. - { econstructor. simplify. rewrite app_nil_r in H. eauto. apply abstract_interp_empty. } - { simplify. rewrite app_assoc in H. rewrite abstract_seq in H. - exploit sem_update'; eauto; simplify. - exploit IHb; eauto; inv_simp. - econstructor; split; eauto. - rewrite abstract_seq. - eapply sem_update2; eauto. - } -Qed.*) - -Lemma sem_update_RBnop : - forall A ctx f st', - @sem A ctx f st' -> sem ctx (update f RBnop) st'. -Proof. auto. Qed. - -Lemma sem_update_Op : - forall A ge sp ist f st' r l o0 o m rs v ps, - @sem A (mk_ctx ist sp ge) f st' -> - eval_predf ps o = true -> - Op.eval_operation ge sp o0 (rs ## l) m = Some v -> - match_states st' (mk_instr_state rs ps m) -> - exists tst, - sem (mk_ctx ist sp ge) (update f (RBop (Some o) o0 l r)) tst - /\ match_states (mk_instr_state (Regmap.set r v rs) ps m) tst. -Proof. - intros. inv H1. inv H. inv H1. inv H3. simplify. - econstructor. simplify. - { constructor; try constructor; intros; try solve [rewrite genmap1; now eauto]. - destruct (Pos.eq_dec x r); subst. - { rewrite map2. specialize (H1 r). inv H1. -(*} - } - destruct st. - econstructor. simplify. - { constructor. - { constructor. intros. destruct (Pos.eq_dec x r); subst. - { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. - { inv H9. eapply gen_list_base; eauto. } - { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } - { inv H. rewrite genmap1; crush. eauto. } } - { constructor; eauto. intros. - destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed.*) Admitted. - -Lemma sem_update : - forall A ge sp st x st' st'' st''' f, - sem (mk_ctx st sp ge) f st' -> - 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. - 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', - RTLBlock.step_instr_list ge sp st bb st' -> - forall tst, - match_states st tst -> - exists tst', sem (mk_ctx tst sp ge) (abstract_sequence empty bb) tst' - /\ match_states st' tst'. -Proof. - induction bb using rev_ind; simplify. - { econstructor. simplify. apply abstract_interp_empty. - inv H. auto. } - { apply rtlblock_trans_correct' in H. simplify. - rewrite abstract_seq. - exploit IHbb; try eassumption; []; simplify. - exploit sem_update. apply H1. symmetry; eassumption. - eauto. simplify. econstructor. split. apply H3. - auto. } -Qed. - -Lemma abstract_execution_correct: - forall bb bb' cfi cfi' ge tge sp st st' tst, - RTLBlock.step_instr_list ge sp st bb st' -> - ge_preserved ge tge -> - schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi') = true -> - match_states st tst -> - exists tst', RTLPar.step_instr_block tge sp tst bb' tst' - /\ match_states st' tst'. -Proof. - intros. - unfold schedule_oracle in *. simplify. unfold empty_trees in H4. - exploit rtlblock_trans_correct; try eassumption; []; simplify. -(*) exploit abstract_execution_correct'; - try solve [eassumption | apply state_lessdef_match_sem; eassumption]. - apply match_states_commut. eauto. inv_simp. - exploit rtlpar_trans_correct; try eassumption; []; inv_simp. - exploit step_instr_block_matches; eauto. apply match_states_commut; eauto. inv_simp. - repeat match goal with | H: match_states _ _ |- _ => inv H end. - do 2 econstructor; eauto. - econstructor; congruence. -Qed.*)Admitted. - -Definition match_prog (prog : RTLBlock.program) (tprog : RTLPar.program) := - match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq prog tprog. - -Inductive match_stackframes: RTLBlock.stackframe -> RTLPar.stackframe -> Prop := -| match_stackframe: - forall f tf res sp pc rs rs' ps ps', - transl_function f = OK tf -> - (forall x, rs !! x = rs' !! x) -> - (forall x, ps !! x = ps' !! x) -> - match_stackframes (Stackframe res f sp pc rs ps) - (Stackframe res tf sp pc rs' ps'). - -Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop := -| match_state: - forall sf f sp pc rs rs' m sf' tf ps ps' - (TRANSL: transl_function f = OK tf) - (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_returnstate: - forall stack stack' v m - (STACKS: list_forall2 match_stackframes stack stack'), - match_states (Returnstate stack v m) - (Returnstate stack' v m) -| match_callstate: - forall stack stack' f tf args m - (TRANSL: transl_fundef f = OK tf) - (STACKS: list_forall2 match_stackframes stack stack'), - match_states (Callstate stack f args m) - (Callstate stack' tf args m). - -Section CORRECTNESS. - - Context (prog: RTLBlock.program) (tprog : RTLPar.program). - Context (TRANSL: match_prog prog tprog). - - Let ge : RTLBlock.genv := Globalenvs.Genv.globalenv prog. - Let tge : RTLPar.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. - Hint Resolve symbols_preserved : rtlgp. - - Lemma function_ptr_translated: - forall (b: Values.block) (f: RTLBlock.fundef), - Genv.find_funct_ptr ge b = Some f -> - exists tf, - Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = Errors.OK tf. - Proof using TRANSL. - intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. - intros (cu & tf & P & Q & R); exists tf; auto. - Qed. - - Lemma functions_translated: - forall (v: Values.val) (f: RTLBlock.fundef), - Genv.find_funct ge v = Some f -> - exists tf, - Genv.find_funct tge v = Some tf /\ transl_fundef f = Errors.OK tf. - Proof using TRANSL. - intros. exploit (Genv.find_funct_match TRANSL); eauto. - intros (cu & tf & P & Q & R); exists tf; auto. - Qed. - - Lemma senv_preserved: - Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). - Proof (Genv.senv_transf_partial TRANSL). - Hint Resolve senv_preserved : rtlgp. - - Lemma sig_transl_function: - forall (f: RTLBlock.fundef) (tf: RTLPar.fundef), - transl_fundef f = OK tf -> - funsig tf = funsig f. - Proof using . - unfold transl_fundef, transf_partial_fundef, transl_function; intros; - repeat destruct_match; crush; - match goal with H: OK _ = OK _ |- _ => inv H end; auto. - Qed. - Hint Resolve sig_transl_function : rtlgp. - - Hint Resolve Val.lessdef_same : rtlgp. - Hint Resolve regs_lessdef_regs : rtlgp. - - Lemma find_function_translated: - forall ros rs rs' f, - (forall x, rs !! x = rs' !! x) -> - find_function ge ros rs = Some f -> - exists tf, find_function tge ros rs' = Some tf - /\ transl_fundef f = OK tf. - Proof using TRANSL. - Ltac ffts := match goal with - | [ H: forall _, Val.lessdef _ _, r: Registers.reg |- _ ] => - specialize (H r); inv H - | [ H: Vundef = ?r, H1: Genv.find_funct _ ?r = Some _ |- _ ] => - rewrite <- H in H1 - | [ H: Genv.find_funct _ Vundef = Some _ |- _] => solve [inv H] - | _ => solve [exploit functions_translated; eauto] - end. - destruct ros; simplify; try rewrite <- H; - [| rewrite symbols_preserved; destruct_match; - try (apply function_ptr_translated); crush ]; - intros; - repeat ffts. - Qed. - - Lemma schedule_oracle_nil: - forall bb cfi, - schedule_oracle {| bb_body := nil; bb_exit := cfi |} bb = true -> - bb_body bb = nil /\ bb_exit bb = cfi. - Proof using . - unfold schedule_oracle, check_control_flow_instr. - simplify; repeat destruct_match; crush. - Qed. - - Lemma schedule_oracle_nil2: - forall cfi, - schedule_oracle {| bb_body := nil; bb_exit := cfi |} - {| bb_body := nil; bb_exit := cfi |} = true. - Proof using . - unfold schedule_oracle, check_control_flow_instr. - simplify; repeat destruct_match; crush. - Qed. - - Lemma eval_op_eq: - forall (sp0 : Values.val) (op : Op.operation) (vl : list Values.val) m, - Op.eval_operation ge sp0 op vl m = Op.eval_operation tge sp0 op vl m. - Proof using TRANSL. - intros. - destruct op; auto; unfold Op.eval_operation, Genv.symbol_address, Op.eval_addressing32; - [| destruct a; unfold Genv.symbol_address ]; - try rewrite symbols_preserved; auto. - Qed. - Hint Resolve eval_op_eq : rtlgp. - - Lemma eval_addressing_eq: - forall sp addr vl, - Op.eval_addressing ge sp addr vl = Op.eval_addressing tge sp addr vl. - Proof using TRANSL. - intros. - destruct addr; - unfold Op.eval_addressing, Op.eval_addressing32; - unfold Genv.symbol_address; - try rewrite symbols_preserved; auto. - Qed. - Hint Resolve eval_addressing_eq : rtlgp. - - Lemma ge_preserved_lem: - ge_preserved ge tge. - Proof using TRANSL. - unfold ge_preserved. - eauto with rtlgp. - Qed. - Hint Resolve ge_preserved_lem : rtlgp. - - Lemma lessdef_regmap_optget: - forall or rs rs', - regs_lessdef rs rs' -> - Val.lessdef (regmap_optget or Vundef rs) (regmap_optget or Vundef rs'). - Proof using. destruct or; crush. Qed. - Hint Resolve lessdef_regmap_optget : rtlgp. - - Lemma regmap_equiv_lessdef: - forall rs rs', - (forall x, rs !! x = rs' !! x) -> - regs_lessdef rs rs'. - Proof using. - intros; unfold regs_lessdef; intros. - rewrite H. apply Val.lessdef_refl. - Qed. - Hint Resolve regmap_equiv_lessdef : rtlgp. - - Lemma int_lessdef: - forall rs rs', - regs_lessdef rs rs' -> - (forall arg v, - rs !! arg = Vint v -> - rs' !! arg = Vint v). - Proof using. intros ? ? H; intros; specialize (H arg); inv H; crush. Qed. - Hint Resolve int_lessdef : rtlgp. - - Ltac semantics_simpl := - match goal with - | [ H: match_states _ _ |- _ ] => - let H2 := fresh "H" in - learn H as H2; inv H2 - | [ H: transl_function ?f = OK _ |- _ ] => - let H2 := fresh "TRANSL" in - learn H as H2; - unfold transl_function in H2; - destruct (check_scheduled_trees - (@fn_code RTLBlock.bb f) - (@fn_code RTLPar.bb (schedule f))) eqn:?; - [| discriminate ]; inv H2 - | [ H: context[check_scheduled_trees] |- _ ] => - let H2 := fresh "CHECK" in - learn H as H2; - eapply check_scheduled_trees_correct in H2; [| solve [eauto] ] - | [ H: schedule_oracle {| bb_body := nil; bb_exit := _ |} ?bb = true |- _ ] => - let H2 := fresh "SCHED" in - learn H as H2; - apply schedule_oracle_nil in H2 - | [ H: find_function _ _ _ = Some _, H2: forall x, ?rs !! x = ?rs' !! x |- _ ] => - learn H; exploit find_function_translated; try apply H2; eauto; inversion 1 - | [ H: Mem.free ?m _ _ _ = Some ?m', H2: Mem.extends ?m ?m'' |- _ ] => - learn H; exploit Mem.free_parallel_extends; eauto; intros - | [ H: Events.eval_builtin_args _ _ _ _ _ _, H2: regs_lessdef ?rs ?rs' |- _ ] => - let H3 := fresh "H" in - learn H; exploit Events.eval_builtin_args_lessdef; [apply H2 | | |]; - eauto with rtlgp; intro H3; learn H3 - | [ H: Events.external_call _ _ _ _ _ _ _ |- _ ] => - let H2 := fresh "H" in - learn H; exploit Events.external_call_mem_extends; - eauto; intro H2; learn H2 - | [ H: exists _, _ |- _ ] => inv H - | _ => progress simplify - end. - - Hint Resolve Events.eval_builtin_args_preserved : rtlgp. - Hint Resolve Events.external_call_symbols_preserved : rtlgp. - Hint Resolve set_res_lessdef : rtlgp. - Hint Resolve set_reg_lessdef : rtlgp. - Hint Resolve Op.eval_condition_lessdef : rtlgp. - - Hint Constructors Events.eval_builtin_arg: barg. - - Lemma eval_builtin_arg_eq: - forall A ge a v1 m1 e1 e2 sp, - (forall x, e1 x = e2 x) -> - @Events.eval_builtin_arg A ge e1 sp m1 a v1 -> - Events.eval_builtin_arg ge e2 sp m1 a v1. -Proof. induction 2; try rewrite H; eauto with barg. Qed. - - Lemma eval_builtin_args_eq: - forall A ge e1 sp m1 e2 al vl1, - (forall x, e1 x = e2 x) -> - @Events.eval_builtin_args A ge e1 sp m1 al vl1 -> - Events.eval_builtin_args ge e2 sp m1 al vl1. - Proof. - induction 2. - - econstructor; split. - - exploit eval_builtin_arg_eq; eauto. intros. - destruct IHlist_forall2 as [| y]. constructor; eauto. - constructor. constructor; auto. - constructor; eauto. - Qed. - - Lemma step_cf_instr_correct: - forall cfi t s s', - step_cf_instr ge s cfi t s' -> - forall r, - match_states s r -> - exists r', step_cf_instr tge r cfi t r' /\ match_states s' r'. - Proof using TRANSL. - induction 1; repeat semantics_simpl; - try solve [repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp)]. - { do 3 (try erewrite match_states_list by eauto; econstructor; eauto with rtlgp). - eapply eval_builtin_args_eq. eapply REG. - eapply Events.eval_builtin_args_preserved. eapply symbols_preserved. - eauto. - intros. - unfold regmap_setres. destruct res. - destruct (Pos.eq_dec x0 x); subst. - repeat rewrite Regmap.gss; auto. - repeat rewrite Regmap.gso; auto. - eapply REG. eapply REG. - } - { repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). - unfold regmap_optget. destruct or. rewrite REG. constructor; eauto. - constructor; eauto. - } - { exploit IHstep_cf_instr; eauto. simplify. - repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). - erewrite eval_predf_pr_equiv; eauto. - } - Qed. - - Theorem transl_step_correct : - forall (S1 : RTLBlock.state) t S2, - RTLBlock.step ge S1 t S2 -> - forall (R1 : RTLPar.state), - match_states S1 R1 -> - exists R2, Smallstep.plus RTLPar.step tge R1 t R2 /\ match_states S2 R2. - Proof. - - induction 1; repeat semantics_simpl. - - { destruct bb; destruct x. - assert (bb_exit = bb_exit0). - { unfold schedule_oracle in *. simplify. - unfold check_control_flow_instr in *. - destruct_match; crush. - } - subst. - - exploit abstract_execution_correct; try eassumption. eapply ge_preserved_lem. - econstructor; eauto. - simplify. destruct x. inv H7. - - exploit step_cf_instr_correct; try eassumption. econstructor; eauto. - simplify. - - econstructor. econstructor. eapply Smallstep.plus_one. econstructor. - 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. - 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. } - { inv STACKS. inv H2. repeat econstructor; eauto. - intros. apply PTree_matches; eauto. } - Qed. - - Lemma transl_initial_states: - forall S, - RTLBlock.initial_state prog S -> - exists R, RTLPar.initial_state tprog R /\ match_states S R. - Proof. - induction 1. - exploit function_ptr_translated; eauto. intros [tf [A B]]. - econstructor; split. - econstructor. apply (Genv.init_mem_transf_partial TRANSL); eauto. - replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto. - symmetry; eapply match_program_main; eauto. - eexact A. - rewrite <- H2. apply sig_transl_function; auto. - constructor. auto. constructor. - Qed. - - Lemma transl_final_states: - forall S R r, - match_states S R -> RTLBlock.final_state S r -> RTLPar.final_state R r. - Proof. - intros. inv H0. inv H. inv STACKS. constructor. - Qed. - - Theorem transf_program_correct: - Smallstep.forward_simulation (RTLBlock.semantics prog) (RTLPar.semantics tprog). - Proof. - eapply Smallstep.forward_simulation_plus. - apply senv_preserved. - eexact transl_initial_states. - eexact transl_final_states. - exact transl_step_correct. - Qed. - -End CORRECTNESS. -#+end_src - -* License - -#+name: license -#+begin_src coq :tangle no -(* - * 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 . - *) -#+end_src -- 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 --- lit/basic-block-generation.org | 2 +- 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 +- 9 files changed, 20 insertions(+), 22 deletions(-) diff --git a/lit/basic-block-generation.org b/lit/basic-block-generation.org index 4a3edc4..b835e61 100644 --- a/lit/basic-block-generation.org +++ b/lit/basic-block-generation.org @@ -29,7 +29,7 @@ Require Import vericert.hls.RTLBlock. #+end_src #+name: rtlblockgen-equalities-insert -#+begin_src coq +#+begin_src coq :comments no <> #+end_src 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 9eeb3845eb466189276fb16e08c41902b430c342 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 24 Mar 2022 10:03:59 +0000 Subject: Rename changelog --- CHANGELOG.org | 65 ----------------------------------------------------------- ChangeLog.org | 65 +++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 65 insertions(+), 65 deletions(-) delete mode 100644 CHANGELOG.org create mode 100644 ChangeLog.org diff --git a/CHANGELOG.org b/CHANGELOG.org deleted file mode 100644 index 88c0953..0000000 --- a/CHANGELOG.org +++ /dev/null @@ -1,65 +0,0 @@ -# -*- fill-column: 80 -*- -#+title: Vericert Changelog -#+author: Yann Herklotz -#+email: git@ymhg.org - -* Unreleased - -** New Features - -- Add ~RTLBlock~, a basic block intermediate language that is based on CompCert's - RTL. -- Add ~RTLPar~, which can execute groups of instructions in parallel. -- Add SDC hyper-block scheduling pass to go from RTLBlock to RTLPar using. -- Add operation chaining support to scheduler. -- Add ~Abstr~ intermediate language for equivalence checking of schedule. -- Add built-in verified SAT solver used for equivalence checking of - hyper-blocks. - -* 2021-10-01 - v1.2.2 - -Mainly fix some documentation and remove any ~Admitted~ theorems, even though -these were in parts of the compiler that were never used. - -* 2021-07-12 - v1.2.1 - -Main release for OOPSLA'21 paper. - -- Add better documentation on how to run Vericert. -- Add =Dockerfile= with instructions on how to get figures of the paper. - -* 2021-04-07 - v1.2.0 - -** New Features - -- Add memory inference capabilities in generated hardware. - -* 2020-12-17 - v1.1.0 - -Add a stable release with all proofs completed. - -* 2020-08-14 - v1.0.1 - -Release a new minor version fixing all proofs and fixing scripts to generate the -badges. - -** Fixes - -- Fix some of the proofs which were not passing. - -* 2020-08-13 - v1.0.0 - -First release of a fully verified version of Vericert with support for the -translation of many C constructs to Verilog. - -** New Features - -- Most int instructions and operators. -- Non-recursive function calls. -- Local arrays, structs and unions of type int. -- Pointer arithmetic with int. - -* 2020-04-03 - v0.1.0 - -This is the first release with working HLS but without any proofs associated -with it. diff --git a/ChangeLog.org b/ChangeLog.org new file mode 100644 index 0000000..88c0953 --- /dev/null +++ b/ChangeLog.org @@ -0,0 +1,65 @@ +# -*- fill-column: 80 -*- +#+title: Vericert Changelog +#+author: Yann Herklotz +#+email: git@ymhg.org + +* Unreleased + +** New Features + +- Add ~RTLBlock~, a basic block intermediate language that is based on CompCert's + RTL. +- Add ~RTLPar~, which can execute groups of instructions in parallel. +- Add SDC hyper-block scheduling pass to go from RTLBlock to RTLPar using. +- Add operation chaining support to scheduler. +- Add ~Abstr~ intermediate language for equivalence checking of schedule. +- Add built-in verified SAT solver used for equivalence checking of + hyper-blocks. + +* 2021-10-01 - v1.2.2 + +Mainly fix some documentation and remove any ~Admitted~ theorems, even though +these were in parts of the compiler that were never used. + +* 2021-07-12 - v1.2.1 + +Main release for OOPSLA'21 paper. + +- Add better documentation on how to run Vericert. +- Add =Dockerfile= with instructions on how to get figures of the paper. + +* 2021-04-07 - v1.2.0 + +** New Features + +- Add memory inference capabilities in generated hardware. + +* 2020-12-17 - v1.1.0 + +Add a stable release with all proofs completed. + +* 2020-08-14 - v1.0.1 + +Release a new minor version fixing all proofs and fixing scripts to generate the +badges. + +** Fixes + +- Fix some of the proofs which were not passing. + +* 2020-08-13 - v1.0.0 + +First release of a fully verified version of Vericert with support for the +translation of many C constructs to Verilog. + +** New Features + +- Most int instructions and operators. +- Non-recursive function calls. +- Local arrays, structs and unions of type int. +- Pointer arithmetic with int. + +* 2020-04-03 - v0.1.0 + +This is the first release with working HLS but without any proofs associated +with it. -- 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 --- docs/basic-block-generation.org | 497 +++++++++ docs/scheduler-languages.org | 689 ++++++++++++ docs/scheduler.org | 2295 +++++++++++++++++++++++++++++++++++++++ lit/basic-block-generation.org | 378 ------- lit/scheduler-languages.org | 689 ------------ lit/scheduler.org | 2295 --------------------------------------- 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 +- 12 files changed, 3499 insertions(+), 3378 deletions(-) create mode 100644 docs/basic-block-generation.org create mode 100644 docs/scheduler-languages.org create mode 100644 docs/scheduler.org delete mode 100644 lit/basic-block-generation.org delete mode 100644 lit/scheduler-languages.org delete mode 100644 lit/scheduler.org diff --git a/docs/basic-block-generation.org b/docs/basic-block-generation.org new file mode 100644 index 0000000..1d78dad --- /dev/null +++ b/docs/basic-block-generation.org @@ -0,0 +1,497 @@ +#+title: Basic Block Generation +#+author: Yann Herklotz +#+email: yann [at] yannherklotz [dot] com + +* RTLBlockgen +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgen.v +:END: + +Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]]. + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblockgen-imports +#+begin_src coq +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. +#+end_src + +#+name: rtlblockgen-equalities-insert +#+begin_src coq :comments no +<> +#+end_src + +#+name: rtlblockgen-main +#+begin_src coq +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 + | 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. +#+end_src + +** Equalities + +#+name: rtlblockgen-equalities +#+begin_src coq :tangle no +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. +#+end_src + +* RTLBlockgenproof +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +** Imports + +#+name: rtlblockgenproof-imports +#+begin_src coq +Require compcert.backend.RTL. +Require Import compcert.common.AST. +Require Import compcert.lib.Maps. + +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLBlockgen. +#+end_src + +** Match states + +The ~match_states~ predicate describes which states are equivalent between the two languages, in this +case ~RTL~ and ~RTLBlock~. + +#+name: rtlblockgenproof-match-states +#+begin_src coq +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). +#+end_src + +** Correctness + +#+name: rtlblockgenproof-correctness +#+begin_src coq +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. +#+end_src + +* Partition +:PROPERTIES: +:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Partition.ml +:END: + +#+begin_src ocaml :comments no :padline no :exports none +<> +#+end_src + +#+name: partition-main +#+begin_src ocaml +open Printf +open Clflags +open Camlcoq +open Datatypes +open Coqlib +open Maps +open AST +open Kildall +open Op +open RTLBlockInstr +open RTLBlock + +(** 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 + let filt = List.filter (fun s -> P.lt n s || P.lt s (P.pred n)) succ in + ((match filt with [] -> [] | _ -> [n]), filt) + +let find_edges c = + PTree.fold (fun l n i -> + let f = find_edge i n in + (List.append (fst f) (fst l), List.append (snd f) (snd l))) c ([], []) + +let prepend_instr i = function + | {bb_body = bb; bb_exit = e} -> {bb_body = (i :: bb); bb_exit = e} + +let translate_inst = function + | RTL.Inop _ -> Some RBnop + | RTL.Iop (op, ls, dst, _) -> Some (RBop (None, op, ls, dst)) + | RTL.Iload (m, addr, ls, dst, _) -> Some (RBload (None, m, addr, ls, dst)) + | RTL.Istore (m, addr, ls, src, _) -> Some (RBstore (None, m, addr, ls, src)) + | _ -> None + +let translate_cfi = function + | RTL.Icall (s, r, ls, dst, n) -> Some (RBcall (s, r, ls, dst, n)) + | RTL.Itailcall (s, r, ls) -> Some (RBtailcall (s, r, ls)) + | RTL.Ibuiltin (e, ls, r, n) -> Some (RBbuiltin (e, ls, r, n)) + | RTL.Icond (c, ls, dst1, dst2) -> Some (RBcond (c, ls, dst1, dst2)) + | RTL.Ijumptable (r, ls) -> Some (RBjumptable (r, ls)) + | RTL.Ireturn r -> Some (RBreturn r) + | _ -> None + +let rec next_bblock_from_RTL is_start e (c : RTL.code) s i = + let succ = List.map (fun i -> (i, PTree.get i c)) (RTL.successors_instr i) in + let trans_inst = (translate_inst i, translate_cfi i) in + 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 = [RBnop]; bb_exit = RBgoto s } + else + 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 = [RBnop]; bb_exit = RBgoto s } + else begin + match next_bblock_from_RTL false e c s' i_n with + | Errors.OK bb -> + Errors.OK (prepend_instr i' bb) + | Errors.Error msg -> Errors.Error msg + end + | _, _ -> + Errors.Error (Errors.msg (coqstring_of_camlstring "next_bblock_from_RTL went wrong.")) + +let rec traverseacc f l c = + match l with + | [] -> Errors.OK c + | x::xs -> + match f x c with + | Errors.Error msg -> Errors.Error msg + | Errors.OK x' -> + match traverseacc f xs x' with + | Errors.Error msg -> Errors.Error msg + | Errors.OK xs' -> Errors.OK xs' + +let rec translate_all edge c s res = + let c_bb, translated = res in + if List.exists (fun x -> P.eq x s) translated then Errors.OK (c_bb, translated) else + (match PTree.get s c with + | None -> Errors.Error (Errors.msg (coqstring_of_camlstring "Could not translate all.")) + | Some i -> + match next_bblock_from_RTL true edge c s i with + | Errors.Error msg -> Errors.Error msg + | Errors.OK {bb_body = bb; bb_exit = e} -> + let succ = List.filter (fun x -> P.lt x s) (successors_instr e) in + (match traverseacc (translate_all edge c) succ (c_bb, s :: translated) with + | Errors.Error msg -> Errors.Error msg + | Errors.OK (c', t') -> + Errors.OK (PTree.set s {bb_body = bb; bb_exit = e} c', t'))) + +(* Partition a function and transform it into RTLBlock. *) +let function_from_RTL f = + let e = find_edges f.RTL.fn_code in + match translate_all e f.RTL.fn_code f.RTL.fn_entrypoint (PTree.empty, []) with + | Errors.Error msg -> Errors.Error msg + | Errors.OK (c, _) -> + Errors.OK { fn_sig = f.RTL.fn_sig; + fn_stacksize = f.RTL.fn_stacksize; + fn_params = f.RTL.fn_params; + fn_entrypoint = f.RTL.fn_entrypoint; + fn_code = c + } + +let partition = function_from_RTL +#+end_src + +* License + +#+name: license +#+begin_src coq :tangle no +(* + * 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 . + *) +#+end_src diff --git a/docs/scheduler-languages.org b/docs/scheduler-languages.org new file mode 100644 index 0000000..cf03b3a --- /dev/null +++ b/docs/scheduler-languages.org @@ -0,0 +1,689 @@ +#+title: Scheduler Languages +#+author: Yann Herklotz +#+email: yann [at] yannherklotz [dot] com + +* RTLBlockInstr +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockInstr.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblockinstr-imports +#+begin_src coq +Require Import Coq.micromega.Lia. + +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.Values. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. +Require Import compcert.verilog.Op. + +Require Import Predicate. +Require Import Vericertlib. +#+end_src + +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]). + +#+name: rtlblockinstr-instr-def +#+begin_src coq +Definition node := positive. + +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. +#+end_src + +** Control-Flow Instruction Definition + +These are the instructions that count as control-flow, and will be placed at the end of the basic +blocks. + +#+name: rtlblockinstr-cf-instr-def +#+begin_src coq +Inductive cf_instr : Type := +| RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr +| RBtailcall : signature -> reg + ident -> list reg -> cf_instr +| RBbuiltin : external_function -> list (builtin_arg reg) -> + builtin_res reg -> node -> cf_instr +| RBcond : condition -> list reg -> node -> node -> cf_instr +| RBjumptable : reg -> list node -> cf_instr +| RBreturn : option reg -> cf_instr +| RBgoto : node -> cf_instr +| RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. +#+end_src + +** Helper functions + +#+name: rtlblockinstr-helpers +#+begin_src coq +Fixpoint successors_instr (i : cf_instr) : list node := + match i with + | RBcall sig ros args res s => s :: nil + | RBtailcall sig ros args => nil + | RBbuiltin ef args res s => s :: nil + | RBcond cond args ifso ifnot => ifso :: ifnot :: nil + | RBjumptable arg tbl => tbl + | RBreturn optarg => nil + | RBgoto n => n :: nil + | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil) + end. + +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) + | RBload p chunk addr args dst => + fold_left Pos.max args (Pos.max dst m) + | RBstore p chunk addr args src => + fold_left Pos.max args (Pos.max src m) + | RBsetpred p' c args p => + 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)) + | RBcall sig (inr id) args res s => + fold_left Pos.max args (Pos.max res m) + | RBtailcall sig (inl r) args => + fold_left Pos.max args (Pos.max r m) + | RBtailcall sig (inr id) args => + 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) + | RBcond cond args ifso ifnot => fold_left Pos.max args m + | RBjumptable arg tbl => Pos.max arg m + | RBreturn None => m + | RBreturn (Some arg) => Pos.max arg m + | RBgoto n => m + | RBpred_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. +#+end_src + +*** 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. + +#+name: rtlblockinstr-instr-state +#+begin_src coq +Record instr_state := mk_instr_state { + is_rs: regset; + is_ps: predset; + is_mem: mem; +}. +#+end_src + +** Top-Level Type Definitions + +#+name: rtlblockinstr-type-def +#+begin_src coq +Section DEFINITION. + + Context {bblock_body: Type}. + + 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_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. + +End DEFINITION. +#+end_src + +** Semantics + +#+name: rtlblockinstr-semantics +#+begin_src coq +Section RELSEM. + + Context {bblock_body : Type}. + + Definition genv := Genv.t (@fundef bblock_body) unit. + + 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_RBnop: + 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 + | 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 + | 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 + | 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. + + 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, + 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') + 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) (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') + | 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) + | 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) + | 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) + 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) + | 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'. + +End RELSEM. +#+end_src + +* RTLBlock +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlock.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlblock-main +#+begin_src coq +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.RTLBlockInstr. + +Definition bb := list instr. + +Definition bblock := @bblock bb. +Definition code := @code bb. +Definition function := @function bb. +Definition fundef := @fundef bb. +Definition program := @program bb. +Definition funsig := @funsig bb. +Definition stackframe := @stackframe 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). + + 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_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). +#+end_src + +* RTLPar +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPar.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpar-main +#+begin_src coq +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.RTLBlockInstr. + +Definition bb := list (list (list instr)). + +Definition bblock := @bblock bb. +Definition code := @code bb. +Definition function := @function bb. +Definition fundef := @fundef bb. +Definition program := @program bb. +Definition funsig := @funsig bb. +Definition stackframe := @stackframe bb. +Definition state := @state bb. +Definition genv := @genv bb. + +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_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 -> bb -> 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: 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 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') + | 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. +#+end_src + +* License + +#+name: license +#+begin_src coq :tangle no +(* + * 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 . + *) +#+end_src diff --git a/docs/scheduler.org b/docs/scheduler.org new file mode 100644 index 0000000..018633c --- /dev/null +++ b/docs/scheduler.org @@ -0,0 +1,2295 @@ +#+title: Basic Block Generation +#+author: Yann Herklotz +#+email: yann [at] yannherklotz [dot] com + +* Scheduler +:PROPERTIES: +:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Schedule.ml +:END: + +#+begin_src ocaml :comments no :padline no :exports none +<> +#+end_src + +#+name: scheduler-main +#+begin_src ocaml +open Printf +open Clflags +open Camlcoq +open Datatypes +open Coqlib +open Maps +open AST +open Kildall +open Op +open RTLBlockInstr +open Predicate +open RTLBlock +open HTL +open Verilog +open HTLgen +open HTLMonad +open HTLMonadExtra + +module SS = Set.Make(P) + +type svtype = + | BBType of int + | VarType of int * int + +type sv = { + sv_type: svtype; + sv_num: int; +} + +let print_sv v = + match v with + | { sv_type = BBType bbi; sv_num = n } -> sprintf "bb%d_%d" bbi n + | { sv_type = VarType (bbi, i); sv_num = n } -> sprintf "var%dn%d_%d" bbi i n + +module G = Graph.Persistent.Digraph.ConcreteLabeled(struct + type t = sv + let compare = compare + let equal = (=) + let hash = Hashtbl.hash +end)(struct + type t = int + let compare = compare + let hash = Hashtbl.hash + let equal = (=) + let default = 0 +end) + +module GDot = Graph.Graphviz.Dot(struct + let graph_attributes _ = [] + let default_vertex_attributes _ = [] + let vertex_name = print_sv + let vertex_attributes _ = [] + let get_subgraph _ = None + let default_edge_attributes _ = [] + let edge_attributes _ = [] + + include G + end) + +module DFG = Graph.Persistent.Digraph.ConcreteLabeled(struct + type t = int * instr + let compare = compare + let equal = (=) + let hash = Hashtbl.hash +end)(struct + type t = int + let compare = compare + let hash = Hashtbl.hash + let equal = (=) + let default = 0 +end) + +module DFGSimp = Graph.Persistent.Graph.Concrete(struct + type t = int * instr + let compare = compare + let equal = (=) + let hash = Hashtbl.hash + end) + +let convert dfg = + DFG.fold_vertex (fun v g -> DFGSimp.add_vertex g v) dfg DFGSimp.empty + |> DFG.fold_edges (fun v1 v2 g -> DFGSimp.add_edge (DFGSimp.add_edge g v1 v2) v2 v1) dfg + +let reg r = sprintf "r%d" (P.to_int r) +let print_pred r = sprintf "p%d" (P.to_int r) + +let print_instr = function + | RBnop -> "" + | RBload (_, _, _, _, r) -> sprintf "load(%s)" (reg r) + | RBstore (_, _, _, _, r) -> sprintf "store(%s)" (reg r) + | RBsetpred (_, _, _, p) -> sprintf "setpred(%s)" (print_pred p) + | RBop (_, op, args, d) -> + (match op, args with + | Omove, _ -> "mov" + | Ointconst n, _ -> sprintf "%s=%ld" (reg d) (camlint_of_coqint n) + | Olongconst n, _ -> sprintf "%s=%LdL" (reg d) (camlint64_of_coqint n) + | Ofloatconst n, _ -> sprintf "%s=%.15F" (reg d) (camlfloat_of_coqfloat n) + | Osingleconst n, _ -> sprintf "%s=%.15Ff" (reg d) (camlfloat_of_coqfloat32 n) + | Oindirectsymbol id, _ -> sprintf "%s=&%s" (reg d) (extern_atom id) + | Ocast8signed, [r1] -> sprintf "%s=int8signed(%s)" (reg d) (reg r1) + | Ocast8unsigned, [r1] -> sprintf "%s=int8unsigned(%s)" (reg d) (reg r1) + | Ocast16signed, [r1] -> sprintf "%s=int16signed(%s)" (reg d) (reg r1) + | Ocast16unsigned, [r1] -> sprintf "%s=int16unsigned(%s)" (reg d) (reg r1) + | Oneg, [r1] -> sprintf "%s=-%s" (reg d) (reg r1) + | Osub, [r1;r2] -> sprintf "%s=%s-%s" (reg d) (reg r1) (reg r2) + | Omul, [r1;r2] -> sprintf "%s=%s*%s" (reg d) (reg r1) (reg r2) + | Omulimm n, [r1] -> sprintf "%s=%s*%ld" (reg d) (reg r1) (camlint_of_coqint n) + | Omulhs, [r1;r2] -> sprintf "%s=mulhs(%s,%s)" (reg d) (reg r1) (reg r2) + | Omulhu, [r1;r2] -> sprintf "%s=mulhu(%s,%s)" (reg d) (reg r1) (reg r2) + | Odiv, [r1;r2] -> sprintf "%s=%s /s %s" (reg d) (reg r1) (reg r2) + | Odivu, [r1;r2] -> sprintf "%s=%s /u %s" (reg d) (reg r1) (reg r2) + | Omod, [r1;r2] -> sprintf "%s=%s %%s %s" (reg d) (reg r1) (reg r2) + | Omodu, [r1;r2] -> sprintf "%s=%s %%u %s" (reg d) (reg r1) (reg r2) + | Oand, [r1;r2] -> sprintf "%s=%s & %s" (reg d) (reg r1) (reg r2) + | Oandimm n, [r1] -> sprintf "%s=%s & %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oor, [r1;r2] -> sprintf "%s=%s | %s" (reg d) (reg r1) (reg r2) + | Oorimm n, [r1] -> sprintf "%s=%s | %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oxor, [r1;r2] -> sprintf "%s=%s ^ %s" (reg d) (reg r1) (reg r2) + | Oxorimm n, [r1] -> sprintf "%s=%s ^ %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Onot, [r1] -> sprintf "%s=not(%s)" (reg d) (reg r1) + | Oshl, [r1;r2] -> sprintf "%s=%s << %s" (reg d) (reg r1) (reg r2) + | Oshlimm n, [r1] -> sprintf "%s=%s << %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshr, [r1;r2] -> sprintf "%s=%s >>s %s" (reg d) (reg r1) (reg r2) + | Oshrimm n, [r1] -> sprintf "%s=%s >>s %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrximm n, [r1] -> sprintf "%s=%s >>x %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshru, [r1;r2] -> sprintf "%s=%s >>u %s" (reg d) (reg r1) (reg r2) + | Oshruimm n, [r1] -> sprintf "%s=%s >>u %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Ororimm n, [r1] -> sprintf "%s=%s ror %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshldimm n, [r1;r2] -> sprintf "%s=(%s, %s) << %ld" (reg d) (reg r1) (reg r2) (camlint_of_coqint n) + | Olea addr, args -> sprintf "%s=addr" (reg d) + | Omakelong, [r1;r2] -> sprintf "%s=makelong(%s,%s)" (reg d) (reg r1) (reg r2) + | Olowlong, [r1] -> sprintf "%s=lowlong(%s)" (reg d) (reg r1) + | Ohighlong, [r1] -> sprintf "%s=highlong(%s)" (reg d) (reg r1) + | Ocast32signed, [r1] -> sprintf "%s=long32signed(%s)" (reg d) (reg r1) + | Ocast32unsigned, [r1] -> sprintf "%s=long32unsigned(%s)" (reg d) (reg r1) + | Onegl, [r1] -> sprintf "%s=-l %s" (reg d) (reg r1) + | Osubl, [r1;r2] -> sprintf "%s=%s -l %s" (reg d) (reg r1) (reg r2) + | Omull, [r1;r2] -> sprintf "%s=%s *l %s" (reg d) (reg r1) (reg r2) + | Omullimm n, [r1] -> sprintf "%s=%s *l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Omullhs, [r1;r2] -> sprintf "%s=mullhs(%s,%s)" (reg d) (reg r1) (reg r2) + | Omullhu, [r1;r2] -> sprintf "%s=mullhu(%s,%s)" (reg d) (reg r1) (reg r2) + | Odivl, [r1;r2] -> sprintf "%s=%s /ls %s" (reg d) (reg r1) (reg r2) + | Odivlu, [r1;r2] -> sprintf "%s=%s /lu %s" (reg d) (reg r1) (reg r2) + | Omodl, [r1;r2] -> sprintf "%s=%s %%ls %s" (reg d) (reg r1) (reg r2) + | Omodlu, [r1;r2] -> sprintf "%s=%s %%lu %s" (reg d) (reg r1) (reg r2) + | Oandl, [r1;r2] -> sprintf "%s=%s &l %s" (reg d) (reg r1) (reg r2) + | Oandlimm n, [r1] -> sprintf "%s=%s &l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Oorl, [r1;r2] -> sprintf "%s=%s |l %s" (reg d) (reg r1) (reg r2) + | Oorlimm n, [r1] -> sprintf "%s=%s |l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Oxorl, [r1;r2] -> sprintf "%s=%s ^l %s" (reg d) (reg r1) (reg r2) + | Oxorlimm n, [r1] -> sprintf "%s=%s ^l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) + | Onotl, [r1] -> sprintf "%s=notl(%s)" (reg d) (reg r1) + | Oshll, [r1;r2] -> sprintf "%s=%s < sprintf "%s=%s < sprintf "%s=%s >>ls %s" (reg d) (reg r1) (reg r2) + | Oshrlimm n, [r1] -> sprintf "%s=%s >>ls %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrxlimm n, [r1] -> sprintf "%s=%s >>lx %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oshrlu, [r1;r2] -> sprintf "%s=%s >>lu %s" (reg d) (reg r1) (reg r2) + | Oshrluimm n, [r1] -> sprintf "%s=%s >>lu %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Ororlimm n, [r1] -> sprintf "%s=%s rorl %ld" (reg d) (reg r1) (camlint_of_coqint n) + | Oleal addr, args -> sprintf "%s=addr" (reg d) + | Onegf, [r1] -> sprintf "%s=negf(%s)" (reg d) (reg r1) + | Oabsf, [r1] -> sprintf "%s=absf(%s)" (reg d) (reg r1) + | Oaddf, [r1;r2] -> sprintf "%s=%s +f %s" (reg d) (reg r1) (reg r2) + | Osubf, [r1;r2] -> sprintf "%s=%s -f %s" (reg d) (reg r1) (reg r2) + | Omulf, [r1;r2] -> sprintf "%s=%s *f %s" (reg d) (reg r1) (reg r2) + | Odivf, [r1;r2] -> sprintf "%s=%s /f %s" (reg d) (reg r1) (reg r2) + | Onegfs, [r1] -> sprintf "%s=negfs(%s)" (reg d) (reg r1) + | Oabsfs, [r1] -> sprintf "%s=absfs(%s)" (reg d) (reg r1) + | Oaddfs, [r1;r2] -> sprintf "%s=%s +fs %s" (reg d) (reg r1) (reg r2) + | Osubfs, [r1;r2] -> sprintf "%s=%s -fs %s" (reg d) (reg r1) (reg r2) + | Omulfs, [r1;r2] -> sprintf "%s=%s *fs %s" (reg d) (reg r1) (reg r2) + | Odivfs, [r1;r2] -> sprintf "%s=%s /fs %s" (reg d) (reg r1) (reg r2) + | Osingleoffloat, [r1] -> sprintf "%s=singleoffloat(%s)" (reg d) (reg r1) + | Ofloatofsingle, [r1] -> sprintf "%s=floatofsingle(%s)" (reg d) (reg r1) + | Ointoffloat, [r1] -> sprintf "%s=intoffloat(%s)" (reg d) (reg r1) + | Ofloatofint, [r1] -> sprintf "%s=floatofint(%s)" (reg d) (reg r1) + | Ointofsingle, [r1] -> sprintf "%s=intofsingle(%s)" (reg d) (reg r1) + | Osingleofint, [r1] -> sprintf "%s=singleofint(%s)" (reg d) (reg r1) + | Olongoffloat, [r1] -> sprintf "%s=longoffloat(%s)" (reg d) (reg r1) + | Ofloatoflong, [r1] -> sprintf "%s=floatoflong(%s)" (reg d) (reg r1) + | Olongofsingle, [r1] -> sprintf "%s=longofsingle(%s)" (reg d) (reg r1) + | Osingleoflong, [r1] -> sprintf "%s=singleoflong(%s)" (reg d) (reg r1) + | Ocmp c, args -> sprintf "%s=cmp" (reg d) + | Osel (c, ty), r1::r2::args -> sprintf "%s=sel" (reg d) + | _, _ -> sprintf "N/a") + +module DFGDot = Graph.Graphviz.Dot(struct + let graph_attributes _ = [] + let default_vertex_attributes _ = [] + let vertex_name = function (i, instr) -> sprintf "\"%d:%s\"" i (print_instr instr) + let vertex_attributes _ = [] + let get_subgraph _ = None + let default_edge_attributes _ = [] + let edge_attributes _ = [] + + include DFG + end) + +module DFGDfs = Graph.Traverse.Dfs(DFG) + +module IMap = Map.Make (struct + type t = int + + let compare = compare +end) + +let gen_vertex instrs i = (i, List.nth instrs i) + +(** The DFG type defines a list of instructions with their data dependencies as [edges], which are + the pairs of integers that represent the index of the instruction in the [nodes]. The edges + always point from left to right. *) + +let print_list f out_chan a = + fprintf out_chan "[ "; + List.iter (fprintf out_chan "%a " f) a; + fprintf out_chan "]" + +let print_tuple out_chan a = + let l, r = a in + fprintf out_chan "(%d,%d)" l r + +(*let print_dfg out_chan dfg = + fprintf out_chan "{ nodes = %a, edges = %a }" + (print_list PrintRTLBlockInstr.print_bblock_body) + dfg.nodes (print_list print_tuple) dfg.edges*) + +let print_dfg = DFGDot.output_graph + +let read_process command = + let buffer_size = 2048 in + let buffer = Buffer.create buffer_size in + let string = Bytes.create buffer_size in + let in_channel = Unix.open_process_in command in + let chars_read = ref 1 in + while !chars_read <> 0 do + chars_read := input in_channel string 0 buffer_size; + Buffer.add_substring buffer (Bytes.to_string string) 0 !chars_read + done; + ignore (Unix.close_process_in in_channel); + Buffer.contents buffer + +let comb_delay = function + | RBnop -> 0 + | RBop (_, op, _, _) -> + (match op with + | Omove -> 0 + | Ointconst _ -> 0 + | Olongconst _ -> 0 + | Ocast8signed -> 0 + | Ocast8unsigned -> 0 + | Ocast16signed -> 0 + | Ocast16unsigned -> 0 + | Oneg -> 0 + | Onot -> 0 + | Oor -> 0 + | Oorimm _ -> 0 + | Oand -> 0 + | Oandimm _ -> 0 + | Oxor -> 0 + | Oxorimm _ -> 0 + | Omul -> 8 + | Omulimm _ -> 8 + | Omulhs -> 8 + | Omulhu -> 8 + | Odiv -> 72 + | Odivu -> 72 + | Omod -> 72 + | Omodu -> 72 + | _ -> 1) + | _ -> 1 + +let pipeline_stages = function + | RBload _ -> 2 + | RBop (_, op, _, _) -> + (match op with + | Odiv -> 32 + | Odivu -> 32 + | Omod -> 32 + | Omodu -> 32 + | _ -> 0) + | _ -> 0 + +(** Add a dependency if it uses a register that was written to previously. *) +let add_dep map i tree dfg curr = + match PTree.get curr tree with + | None -> dfg + | Some ip -> + let ipv = (List.nth map ip) in + 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. + + This function only gathers the RAW constraints, and will therefore only be active for operations + that modify registers, which is this case only affects loads and operations. *) +let accumulate_RAW_deps map dfg curr = + let i, dst_map, graph = dfg in + let acc_dep_instruction rs dst = + ( i + 1, + PTree.set dst i dst_map, + List.fold_left (add_dep map i dst_map) graph rs + ) + in + let acc_dep_instruction_nodst rs = + ( i + 1, + dst_map, + List.fold_left (add_dep map i dst_map) graph rs) + in + 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) + +(** Finds the next write to the [dst] register. This is a small optimisation so that only one + dependency is generated for a data dependency. *) +let rec find_next_dst_write i dst i' curr = + let check_dst dst' curr' = + if dst = dst' then Some (i, i') + else find_next_dst_write i dst (i' + 1) curr' + in + match curr with + | [] -> None + | RBop (_, _, _, dst') :: curr' -> check_dst dst' curr' + | RBload (_, _, _, _, dst') :: curr' -> check_dst dst' curr' + | _ :: curr' -> find_next_dst_write i dst (i' + 1) curr' + +let rec find_all_next_dst_read i dst i' curr = + let check_dst rs curr' = + if List.exists (fun x -> x = dst) rs + then (i, i') :: find_all_next_dst_read i dst (i' + 1) curr' + else find_all_next_dst_read i dst (i' + 1) curr' + in + match curr with + | [] -> [] + | RBop (_, _, rs, _) :: curr' -> check_dst rs curr' + | RBload (_, _, _, rs, _) :: curr' -> check_dst rs curr' + | RBstore (_, _, _, rs, src) :: curr' -> check_dst (src :: rs) curr' + | RBnop :: curr' -> find_all_next_dst_read i dst (i' + 1) curr' + | RBsetpred (_, _, rs, _) :: curr' -> check_dst rs curr' + +let drop i lst = + let rec drop' i' lst' = + match lst' with + | _ :: ls -> if i' = i then ls else drop' (i' + 1) ls + | [] -> [] + in + if i = 0 then lst else drop' 1 lst + +let take i lst = + let rec take' i' lst' = + match lst' with + | l :: ls -> if i' = i then [ l ] else l :: take' (i' + 1) ls + | [] -> [] + in + if i = 0 then [] else take' 1 lst + +let rec next_store i = function + | [] -> None + | RBstore (_, _, _, _, _) :: _ -> Some i + | _ :: rst -> next_store (i + 1) rst + +let rec next_load i = function + | [] -> None + | RBload (_, _, _, _, _) :: _ -> Some i + | _ :: rst -> next_load (i + 1) rst + +let accumulate_RAW_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBload (_, _, _, _, _) -> ( + 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) ) + | _ -> dfg + +let accumulate_WAR_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBstore (_, _, _, _, _) -> ( + match next_load 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) ) + | _ -> dfg + +let accumulate_WAW_mem_deps instrs dfg curri = + let i, curr = curri in + match curr with + | 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)) + | _ -> dfg + +(** Predicate dependencies. *) + +let rec in_predicate p p' = + match p' with + | Plit p'' -> P.to_int p = P.to_int (snd p'') + | Pand (p1, p2) -> in_predicate p p1 || in_predicate p p2 + | Por (p1, p2) -> in_predicate p p1 || in_predicate p p2 + | Ptrue -> false + | Pfalse -> false + +let get_predicate = function + | RBop (p, _, _, _) -> p + | RBload (p, _, _, _, _) -> p + | RBstore (p, _, _, _, _) -> p + | RBsetpred (p, _, _, _) -> p + | _ -> None + +let rec next_setpred p i = function + | [] -> None + | RBsetpred (_, _, _, p') :: rst -> + if in_predicate p' p then + Some i + else + next_setpred p (i + 1) rst + | _ :: rst -> next_setpred p (i + 1) rst + +let rec next_preduse p i instr= + let next p' rst = + if in_predicate p p' then + Some i + else + next_preduse p (i + 1) rst + in + match instr with + | [] -> None + | 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 = + let i, curr = curri in + match get_predicate curr with + | Some p -> ( + match next_setpred p 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) ) + | _ -> dfg + +let accumulate_WAR_pred_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBsetpred (_, _, _, p) -> ( + match next_preduse p 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) ) + | _ -> dfg + +let accumulate_WAW_pred_deps instrs dfg curri = + let i, curr = curri in + match curr with + | RBsetpred (_, _, _, p) -> ( + match next_setpred (Plit (true, p)) 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) ) + | _ -> dfg + +(** This function calculates the WAW dependencies, which happen when two writes are ordered one + after another and therefore have to be kept in that order. This accumulation might be redundant + if register renaming is done before hand, because then these dependencies can be avoided. *) +let accumulate_WAW_deps instrs dfg curri = + let i, curr = curri in + let dst_dep dst = + match find_next_dst_write i dst (i + 1) (drop (i + 1) instrs) with + | Some (a, b) -> DFG.add_edge dfg (gen_vertex instrs a) (gen_vertex instrs b) + | _ -> dfg + in + match curr with + | RBop (_, _, _, dst) -> dst_dep dst + | RBload (_, _, _, _, dst) -> dst_dep dst + | RBstore (_, _, _, _, _) -> ( + match next_store (i + 1) (drop (i + 1) instrs) with + | None -> dfg + | Some i' -> DFG.add_edge dfg (gen_vertex instrs i) (gen_vertex instrs i') ) + | _ -> dfg + +let accumulate_WAR_deps instrs dfg curri = + let i, curr = curri in + let dst_dep dst = + let dep_list = find_all_next_dst_read i dst 0 (take i instrs |> List.rev) + |> List.map (function (d, d') -> (i - d' - 1, d)) + in + List.fold_left (fun g -> + function (d, d') -> DFG.add_edge g (gen_vertex instrs d) (gen_vertex instrs d')) dfg dep_list + in + match curr with + | RBop (_, _, _, dst) -> dst_dep dst + | RBload (_, _, _, _, dst) -> dst_dep dst + | _ -> dfg + +let assigned_vars vars = function + | RBnop -> vars + | RBop (_, _, _, dst) -> dst :: vars + | RBload (_, _, _, _, dst) -> dst :: vars + | RBstore (_, _, _, _, _) -> vars + | RBsetpred (_, _, _, _) -> vars + +let get_pred = function + | RBnop -> None + | RBop (op, _, _, _) -> op + | RBload (op, _, _, _, _) -> op + | RBstore (op, _, _, _, _) -> op + | RBsetpred (_, _, _, _) -> None + +let independant_pred p p' = + match sat_pred_simple (Pand (p, p')) with + | None -> true + | _ -> false + +let check_dependent op1 op2 = + match op1, op2 with + | Some p, Some p' -> not (independant_pred p p') + | _, _ -> true + +let remove_unnecessary_deps graph = + let is_dependent v1 v2 g = + let (_, instr1) = v1 in + let (_, instr2) = v2 in + if check_dependent (get_pred instr1) (get_pred instr2) + then g + else DFG.remove_edge g v1 v2 + in + DFG.fold_edges is_dependent graph graph + +(** All the nodes in the DFG have to come after the source of the basic block, and should terminate + before the sink of the basic block. After that, there should be constraints for data + dependencies between nodes. *) +let gather_bb_constraints debug bb = + let ibody = List.mapi (fun i a -> (i, a)) bb.bb_body in + let dfg = List.fold_left (fun dfg v -> DFG.add_vertex dfg v) DFG.empty ibody in + let _, _, dfg' = + List.fold_left (accumulate_RAW_deps ibody) + (0, PTree.empty, dfg) + bb.bb_body + in + let dfg'' = List.fold_left (fun dfg f -> List.fold_left (f bb.bb_body) dfg ibody) dfg' + [ accumulate_WAW_deps; + accumulate_WAR_deps; + accumulate_RAW_mem_deps; + accumulate_WAR_mem_deps; + accumulate_WAW_mem_deps; + accumulate_RAW_pred_deps; + accumulate_WAR_pred_deps; + accumulate_WAW_pred_deps + ] + in + let dfg''' = remove_unnecessary_deps dfg'' in + (List.length bb.bb_body, dfg''', successors_instr bb.bb_exit) + +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 + then G.add_edge_e g (encode_bb n 0, 0, encode_var n (fst v1) 0) + else g) |> + (fun g' -> + if DFG.out_degree dfg v1 = 0 + 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, 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 = + List.fold_left (fun g n' -> + G.add_edge_e g (encode_bb n 1, -1, encode_bb n' 0) + ) constr succs + +module BFDFG = Graph.Path.BellmanFord(DFG)(struct + include DFG + type t = int + let weight = DFG.E.label + let compare = compare + let add = (+) + let zero = 0 + end) + +module TopoDFG = Graph.Topological.Make(DFG) + +let negate_graph constr = + DFG.fold_edges_e (function (v1, e, v2) -> fun g -> + DFG.add_edge_e g (v1, -e, v2) + ) constr DFG.empty + +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 + |> 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 maxf))) - 1), + encode_var n (fst v') 0) + ) constr (DFG.fold_vertex (fun v l -> + 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 = + | Mem + | SDiv + | UDiv + +type resources = { + res_mem: DFG.V.t list; + res_udiv: DFG.V.t list; + res_sdiv: DFG.V.t list; +} + +let find_resource = function + | RBload _ -> Some Mem + | RBstore _ -> Some Mem + | RBop (_, op, _, _) -> + ( match op with + | Odiv -> Some SDiv + | Odivu -> Some UDiv + | Omod -> Some SDiv + | Omodu -> Some UDiv + | _ -> None ) + | _ -> None + +let add_resource_constr n dfg constr = + let res = TopoDFG.fold (function (i, instr) -> + function {res_mem = ml; res_sdiv = sdl; res_udiv = udl} as r -> + match find_resource instr with + | Some SDiv -> {r with res_sdiv = (i, instr) :: sdl} + | Some UDiv -> {r with res_udiv = (i, instr) :: udl} + | Some Mem -> {r with res_mem = (i, instr) :: ml} + | None -> r + ) dfg {res_mem = []; res_sdiv = []; res_udiv = []} + in + let get_constraints l g = List.fold_left (fun gv v' -> + match gv with + | (g, None) -> (g, Some v') + | (g, Some v) -> + (G.add_edge_e g (encode_var n (fst v) 0, -1, encode_var n (fst v') 0), Some v') + ) (g, None) l |> fst + in + get_constraints (List.rev res.res_mem) constr + |> get_constraints (List.rev res.res_udiv) + |> get_constraints (List.rev res.res_sdiv) + +let gather_cfg_constraints c constr curr = + let (n, dfg) = curr in + match PTree.get (P.of_int n) c with + | 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 + |> List.filter (fun n' -> n' < n)) + |> add_cycle_constr 8 n dfg + |> add_resource_constr n dfg + +let rec intersperse s = function + | [] -> [] + | [ a ] -> [ a ] + | x :: xs -> x :: s :: intersperse s xs + +let print_objective constr = + let vars = G.fold_vertex (fun v1 l -> + match v1 with + | { sv_type = VarType _; sv_num = 0 } -> print_sv v1 :: l + | _ -> l + ) constr [] + in + "min: " ^ String.concat "" (intersperse " + " vars) ^ ";\n" + +let print_lp constr = + print_objective constr ^ + (G.fold_edges_e (function (e1, w, e2) -> fun s -> + s ^ sprintf "%s - %s <= %d;\n" (print_sv e1) (print_sv e2) w + ) constr "" |> + G.fold_vertex (fun v1 s -> + s ^ sprintf "int %s;\n" (print_sv v1) + ) constr) + +let update_schedule v = function Some l -> Some (v :: l) | None -> Some [ v ] + +let parse_soln (tree, bbtree) s = + let r = Str.regexp "var\\([0-9]+\\)n\\([0-9]+\\)_0[ ]+\\([0-9]+\\)" in + 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 (fn, oc) = Filename.open_temp_file "vericert_" "_lp_solve" in + fprintf oc "%s\n" (print_lp constr); + close_out oc; + + let res = Str.split (Str.regexp_string "\n") (read_process ("lp_solve " ^ fn)) + |> drop 3 + |> List.fold_left parse_soln (IMap.empty, 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 + List.fold_left (fun g v -> + List.fold_left (fun g' v' -> + let edges = DFG.find_all_edges dfg v v' in + List.fold_left DFG.add_edge_e g' edges + ) g l + ) dfg' l + +let rec all_successors dfg v = + List.concat (List.fold_left (fun l v -> + all_successors dfg v :: l + ) [] (DFG.succ dfg v)) + +let order_instr dfg = + DFG.fold_vertex (fun v li -> + if DFG.in_degree dfg v = 0 + then (List.map snd (v :: all_successors dfg v)) :: li + else li + ) dfg [] + +let combine_bb_schedule schedule s = + let i, st = s in + IMap.update st (update_schedule i) schedule + +(**let add_el dfg i l = + List.*) + +let check_in el = + List.exists (List.exists ((=) el)) + +let all_dfs dfg = + let roots = DFG.fold_vertex (fun v li -> + if DFG.in_degree dfg v = 0 then v :: li else li + ) dfg [] in + let dfg' = DFG.fold_edges (fun v1 v2 g -> DFG.add_edge g v2 v1) dfg dfg in + List.fold_left (fun a el -> + if check_in el a then a else + (DFGDfs.fold_component (fun v l -> v :: l) [] dfg' el) :: a) [] roots + +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 = + 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 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 + |> 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 -> + all_dfs x + |> List.map (subgraph x) + |> List.map (fun y -> + TopoDFG.fold (fun i l -> snd i :: l) y [] + |> List.rev))) body2 + (*|> (fun x -> TopoDFG.fold (fun i l -> snd i :: l) x []) + |> List.rev) body2*) + in + { bb_body = final_body2; + bb_exit = ctrl_flow + } + in + PTree.map f c + +let schedule entry (c : RTLBlock.bb RTLBlockInstr.code) = + let debug = true in + let transf_graph (_, dfg, _) = dfg in + let c' = PTree.map1 (fun x -> gather_bb_constraints false x |> transf_graph) c in + (*let _ = if debug then PTree.map (fun r o -> printf "##### %d #####\n%a\n\n" (P.to_int r) print_dfg o) c' else PTree.empty in*) + let cgraph = PTree.elements c' + |> List.map (function (x, y) -> (P.to_int x, y)) + |> List.fold_left (gather_cfg_constraints c) G.empty + in + let graph = open_out "constr_graph.dot" in + fprintf graph "%a\n" GDot.output_graph cgraph; + close_out graph; + let schedule' = solve_constraints cgraph in + (**IMap.iter (fun a b -> printf "##### %d #####\n%a\n\n" a (print_list print_tuple) b) schedule';*) + (**printf "Schedule: %a\n" (fun a x -> IMap.iter (fun d -> fprintf a "%d: %a\n" d (print_list print_tuple)) x) schedule';*) + transf_rtlpar c c' schedule' + +let rec find_reachable_states c e = + match PTree.get e c with + | Some { bb_exit = ex; _ } -> + e :: List.fold_left (fun x a -> List.concat [x; find_reachable_states c a]) [] + (successors_instr ex |> List.filter (fun x -> P.lt x e)) + | None -> assert false + +let add_to_tree c nt i = + match PTree.get i c with + | Some p -> PTree.set i p nt + | None -> assert false + +let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = + let scheduled = schedule f.fn_entrypoint f.fn_code in + let reachable = find_reachable_states scheduled f.fn_entrypoint + |> List.to_seq |> SS.of_seq |> SS.to_seq |> List.of_seq in + { fn_sig = f.fn_sig; + fn_params = f.fn_params; + fn_stacksize = f.fn_stacksize; + fn_code = scheduled (*List.fold_left (add_to_tree scheduled) PTree.empty reachable*); + fn_entrypoint = f.fn_entrypoint + } +#+end_src + +* RTLPargen +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargen.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpargen-main +#+begin_src coq +Require Import compcert.backend.Registers. +Require Import compcert.common.AST. +Require Import compcert.common.Globalenvs. +Require Import compcert.common.Memory. +Require Import compcert.common.Values. +Require Import compcert.lib.Floats. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. +Require compcert.verilog.Op. + +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLPar. +Require Import vericert.hls.RTLBlockInstr. +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. + +(** ** 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) + end. + +Fixpoint replicate {A} (n: nat) (l: A) := + match n with + | O => nil + | S n => l :: replicate n l + end. + +Definition merge''' x y := + match x, y with + | Some p1, Some p2 => Some (Pand p1 p2) + | Some p, None | None, Some p => Some p + | None, None => None + end. + +Definition merge'' x := + match x with + | ((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 := + match pa, pf with + | (p, a), (p', f) => (merge''' p p', f a) + end. + +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. + +(*map (fun x => (fst x, Econs (snd x) Enil)) pel*) +Definition merge' (pel: pred_expr) (tpel: predicated expression_list) := + predicated_map (uncurry Econs) (predicated_prod pel tpel). + +Fixpoint merge (pel: list pred_expr): predicated expression_list := + match pel with + | nil => NE.singleton (T, Enil) + | a :: b => merge' a (merge b) + end. + +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 := + 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 := + 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 := + NE.map (fun x => (fst x, (snd x) pa pb pc)) pf. + +Definition predicated_from_opt {A: Type} (p: option pred_op) (a: A) := + match p with + | Some p' => NE.singleton (p', a) + | None => NE.singleton (T, a) + end. + +#[local] Open Scope non_empty_scope. +#[local] Open Scope pred_op. + +Fixpoint NEfold_left {A B} (f: A -> B -> A) (l: NE.non_empty B) (a: A) : A := + match l with + | NE.singleton a' => f a a' + | a' ::| b => NEfold_left f b (f a a') + end. + +Fixpoint NEapp {A} (l m: NE.non_empty A) := + match l with + | NE.singleton a => a ::| m + | a ::| b => a ::| NEapp b m + end. + +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. + +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 + +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 + | RBop p op rl r => + f # (Reg r) <- + (app_predicated p + (f # (Reg r)) + (map_predicated (pred_ret (Eop op)) (merge (list_translation rl f)))) + | RBload p chunk addr rl r => + f # (Reg r) <- + (app_predicated p + (f # (Reg r)) + (map_predicated + (map_predicated (pred_ret (Eload chunk addr)) (merge (list_translation rl f))) + (f # Mem))) + | RBstore p chunk addr rl r => + f # Mem <- + (app_predicated p + (f # Mem) + (map_predicated + (map_predicated + (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)))) + 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. + +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. + +(** 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. *) + +Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := + match bb with + | nil => + match bbt with + | nil => true + | _ => false + end + | _ => true + end. + +Definition schedule_oracle (bb: RTLBlock.bblock) (bbt: RTLPar.bblock) : bool := + check (abstract_sequence empty (bb_body bb)) + (abstract_sequence empty (concat (concat (bb_body bbt)))) && + check_control_flow_instr (bb_exit bb) (bb_exit bbt) && + empty_trees (bb_body bb) (bb_body bbt). + +Definition check_scheduled_trees := beq2 schedule_oracle. + +Ltac solve_scheduled_trees_correct := + intros; unfold check_scheduled_trees in *; + match goal with + | [ H: context[beq2 _ _ _], x: positive |- _ ] => + rewrite beq2_correct in H; specialize (H x) + end; repeat destruct_match; crush. + +Lemma check_scheduled_trees_correct: + forall f1 f2 x y1, + check_scheduled_trees f1 f2 = true -> + PTree.get x f1 = Some y1 -> + exists y2, PTree.get x f2 = Some y2 /\ schedule_oracle y1 y2 = true. +Proof. solve_scheduled_trees_correct; eexists; crush. Qed. + +Lemma check_scheduled_trees_correct2: + forall f1 f2 x, + check_scheduled_trees f1 f2 = true -> + PTree.get x f1 = None -> + PTree.get x f2 = None. +Proof. solve_scheduled_trees_correct. Qed. + +(** ** Top-level Functions *) + +Parameter schedule : RTLBlock.function -> 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) + f.(fn_params) + f.(fn_stacksize) + tfcode + f.(fn_entrypoint)) + else + 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. +#+end_src +* RTLPargenproof +:PROPERTIES: +:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargenproof.v +:END: + +#+begin_src coq :comments no :padline no :exports none +<> +#+end_src + +#+name: rtlpargenproof-imports +#+begin_src coq +Require Import compcert.backend.Registers. +Require Import compcert.common.AST. +Require Import compcert.common.Errors. +Require Import compcert.common.Linking. +Require Import compcert.common.Globalenvs. +Require Import compcert.common.Memory. +Require Import compcert.common.Values. +Require Import compcert.lib.Maps. + +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLPar. +Require Import vericert.hls.RTLBlockInstr. +Require Import vericert.hls.RTLPargen. +Require Import vericert.hls.Predicate. +Require Import vericert.hls.Abstr. + +#[local] Open Scope positive. +#[local] Open Scope forest. +#[local] Open Scope pred_op. +#+end_src + +** RTLBlock to abstract translation + +Correctness of translation from RTLBlock to the abstract interpretation language. + +#+name: rtlpargenproof-main +#+begin_src coq +(*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. + +Inductive state_lessdef : instr_state -> instr_state -> Prop := + state_lessdef_intro : + forall rs1 rs2 m1, + (forall x, rs1 !! x = rs2 !! x) -> + state_lessdef (mk_instr_state rs1 m1) (mk_instr_state rs2 m1). + +Ltac inv_simp := + repeat match goal with + | H: exists _, _ |- _ => inv H + end; simplify. + +*) + +Definition check_dest i r' := + match i with + | RBop p op rl r => (r =? r')%positive + | RBload p chunk addr rl r => (r =? r')%positive + | _ => false + end. + +Lemma check_dest_dec i r : {check_dest i r = true} + {check_dest i r = false}. +Proof. destruct (check_dest i r); tauto. Qed. + +Fixpoint check_dest_l l r := + match l with + | nil => false + | a :: b => check_dest a r || check_dest_l b r + end. + +Lemma check_dest_l_forall : + forall l r, + check_dest_l l r = false -> + Forall (fun x => check_dest x r = false) l. +Proof. induction l; crush. Qed. + +Lemma check_dest_l_dec i r : {check_dest_l i r = true} + {check_dest_l i r = false}. +Proof. destruct (check_dest_l i r); tauto. Qed. + +Lemma check_dest_update : + forall f i r, + check_dest i r = false -> + (update f i) # (Reg r) = f # (Reg r). +Proof. + destruct i; crush; try apply Pos.eqb_neq in H; apply genmap1; crush. +Qed. + +Lemma check_dest_l_forall2 : + forall l r, + Forall (fun x => check_dest x r = false) l -> + check_dest_l l r = false. +Proof. + induction l; crush. + inv H. apply orb_false_intro; crush. +Qed. + +Lemma check_dest_l_ex2 : + forall l r, + (exists a, In a l /\ check_dest a r = true) -> + check_dest_l l r = true. +Proof. + induction l; crush. + specialize (IHl r). inv H. + apply orb_true_intro; crush. + apply orb_true_intro; crush. + right. apply IHl. exists x. auto. +Qed. + +Lemma check_list_l_false : + forall l x r, + check_dest_l (l ++ x :: nil) r = false -> + check_dest_l l r = false /\ check_dest x r = false. +Proof. + simplify. + apply check_dest_l_forall in H. apply Forall_app in H. + simplify. apply check_dest_l_forall2; auto. + apply check_dest_l_forall in H. apply Forall_app in H. + simplify. inv H1. auto. +Qed. + +Lemma check_dest_l_ex : + forall l r, + check_dest_l l r = true -> + exists a, In a l /\ check_dest a r = true. +Proof. + induction l; crush. + destruct (check_dest a r) eqn:?; try solve [econstructor; crush]. + simplify. + exploit IHl. apply H. simplify. econstructor. simplify. right. eassumption. + auto. +Qed. + +Lemma check_list_l_true : + forall l x r, + check_dest_l (l ++ x :: nil) r = true -> + check_dest_l l r = true \/ check_dest x r = true. +Proof. + simplify. + apply check_dest_l_ex in H; simplify. + apply in_app_or in H. inv H. left. + apply check_dest_l_ex2. exists x0. auto. + inv H0; auto. +Qed. + +Lemma check_dest_l_dec2 l r : + {Forall (fun x => check_dest x r = false) l} + + {exists a, In a l /\ check_dest a r = true}. +Proof. + destruct (check_dest_l_dec l r); [right | left]; + auto using check_dest_l_ex, check_dest_l_forall. +Qed. + +Lemma abstr_comp : + forall l i f x x0, + abstract_sequence f (l ++ i :: nil) = x -> + abstract_sequence f l = x0 -> + x = update x0 i. +Proof. induction l; intros; crush; eapply IHl; eauto. Qed. + +(* + +Lemma gen_list_base: + forall FF ge sp l rs exps st1, + (forall x, @sem_value FF ge sp st1 (exps # (Reg x)) (rs !! x)) -> + sem_val_list ge sp st1 (list_translation l exps) rs ## l. +Proof. + induction l. + intros. simpl. constructor. + intros. simpl. eapply Scons; eauto. +Qed. + +Lemma check_dest_update2 : + forall f r rl op p, + (update f (RBop p op rl r)) # (Reg r) = Eop op (list_translation rl f) (f # Mem). +Proof. crush; rewrite map2; auto. Qed. + +Lemma check_dest_update3 : + forall f r rl p addr chunk, + (update f (RBload p chunk addr rl r)) # (Reg r) = Eload chunk addr (list_translation rl f) (f # Mem). +Proof. crush; rewrite map2; auto. Qed. + +Lemma abstract_seq_correct_aux: + forall FF ge sp i st1 st2 st3 f, + @step_instr FF ge sp st3 i st2 -> + sem ge sp st1 f st3 -> + sem ge sp st1 (update f i) st2. +Proof. + intros; inv H; simplify. + { simplify; eauto. } (*apply match_states_refl. }*) + { inv H0. inv H6. destruct st1. econstructor. simplify. + constructor. intros. + destruct (resource_eq (Reg res) (Reg x)). inv e. + rewrite map2. econstructor. eassumption. apply gen_list_base; eauto. + rewrite Regmap.gss. eauto. + assert (res <> x). { unfold not in *. intros. apply n. rewrite H0. auto. } + rewrite Regmap.gso by auto. + rewrite genmap1 by auto. auto. + + rewrite genmap1; crush. } + { inv H0. inv H7. constructor. constructor. intros. + destruct (Pos.eq_dec dst x); subst. + rewrite map2. econstructor; eauto. + apply gen_list_base. auto. rewrite Regmap.gss. auto. + rewrite genmap1. rewrite Regmap.gso by auto. auto. + unfold not in *; intros. inv H0. auto. + rewrite genmap1; crush. + } + { inv H0. inv H7. constructor. constructor; intros. + rewrite genmap1; crush. + rewrite map2. econstructor; eauto. + apply gen_list_base; auto. + } +Qed. + +Lemma regmap_list_equiv : + forall A (rs1: Regmap.t A) rs2, + (forall x, rs1 !! x = rs2 !! x) -> + forall rl, rs1##rl = rs2##rl. +Proof. induction rl; crush. Qed. + +Lemma sem_update_Op : + forall A ge sp st f st' r l o0 o m rs v, + @sem A ge sp st f st' -> + Op.eval_operation ge sp o0 rs ## l m = Some v -> + match_states st' (mk_instr_state rs m) -> + exists tst, + sem ge sp st (update f (RBop o o0 l r)) tst /\ match_states (mk_instr_state (Regmap.set r v rs) m) tst. +Proof. + intros. inv H1. simplify. + destruct st. + econstructor. simplify. + { constructor. + { constructor. intros. destruct (Pos.eq_dec x r); subst. + { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. + { inv H9. eapply gen_list_base; eauto. } + { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } + { inv H. rewrite genmap1; crush. eauto. } } + { constructor; eauto. intros. + destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed. + +Lemma sem_update_load : + forall A ge sp st f st' r o m a l m0 rs v a0, + @sem A ge sp st f st' -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.loadv m m0 a0 = Some v -> + match_states st' (mk_instr_state rs m0) -> + exists tst : instr_state, + sem ge sp st (update f (RBload o m a l r)) tst + /\ match_states (mk_instr_state (Regmap.set r v rs) m0) tst. +Proof. + intros. inv H2. pose proof H. inv H. inv H9. + destruct st. + econstructor; simplify. + { constructor. + { constructor. intros. + destruct (Pos.eq_dec x r); subst. + { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. + rewrite <- H6. eauto. + instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. auto. } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. } } + { rewrite genmap1; crush. eauto. } } + { constructor; auto; intros. destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed. + +Lemma sem_update_store : + forall A ge sp a0 m a l r o f st m' rs m0 st', + @sem A ge sp st f st' -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.storev m m0 a0 rs !! r = Some m' -> + match_states st' (mk_instr_state rs m0) -> + exists tst, sem ge sp st (update f (RBstore o m a l r)) tst + /\ match_states (mk_instr_state rs m') tst. +Proof. + intros. inv H2. pose proof H. inv H. inv H9. + destruct st. + econstructor; simplify. + { econstructor. + { econstructor; intros. rewrite genmap1; crush. } + { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. rewrite <- H6. + eauto. specialize (H6 r). rewrite H6. eauto. } } + { econstructor; eauto. } +Qed. + +Lemma sem_update : + forall A ge sp st x st' st'' st''' f, + sem ge sp st f st' -> + match_states st' st''' -> + @step_instr A ge sp st''' x st'' -> + exists tst, sem ge sp st (update f x) tst /\ match_states st'' tst. +Proof. + intros. destruct x; inv H1. + { econstructor. split. + apply sem_update_RBnop. eassumption. + apply match_states_commut. auto. } + { eapply sem_update_Op; eauto. } + { eapply sem_update_load; eauto. } + { eapply sem_update_store; eauto. } +Qed. + +Lemma sem_update2_Op : + forall A ge sp st f r l o0 o m rs v, + @sem A ge sp st f (mk_instr_state rs m) -> + Op.eval_operation ge sp o0 rs ## l m = Some v -> + sem ge sp st (update f (RBop o o0 l r)) (mk_instr_state (Regmap.set r v rs) m). +Proof. + intros. destruct st. constructor. + inv H. inv H6. + { constructor; intros. simplify. + destruct (Pos.eq_dec r x); subst. + { rewrite map2. econstructor. eauto. + apply gen_list_base. eauto. + rewrite Regmap.gss. auto. } + { rewrite genmap1; crush. rewrite Regmap.gso; auto. } } + { simplify. rewrite genmap1; crush. inv H. eauto. } +Qed. + +Lemma sem_update2_load : + forall A ge sp st f r o m a l m0 rs v a0, + @sem A ge sp st f (mk_instr_state rs m0) -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.loadv m m0 a0 = Some v -> + sem ge sp st (update f (RBload o m a l r)) (mk_instr_state (Regmap.set r v rs) m0). +Proof. + intros. simplify. inv H. inv H7. constructor. + { constructor; intros. destruct (Pos.eq_dec r x); subst. + { rewrite map2. rewrite Regmap.gss. econstructor; eauto. + apply gen_list_base; eauto. } + { rewrite genmap1; crush. rewrite Regmap.gso; eauto. } + } + { simplify. rewrite genmap1; crush. } +Qed. + +Lemma sem_update2_store : + forall A ge sp a0 m a l r o f st m' rs m0, + @sem A ge sp st f (mk_instr_state rs m0) -> + Op.eval_addressing ge sp a rs ## l = Some a0 -> + Mem.storev m m0 a0 rs !! r = Some m' -> + sem ge sp st (update f (RBstore o m a l r)) (mk_instr_state rs m'). +Proof. + intros. simplify. inv H. inv H7. constructor; simplify. + { econstructor; intros. rewrite genmap1; crush. } + { rewrite map2. econstructor; eauto. apply gen_list_base; eauto. } +Qed. + +Lemma sem_update2 : + forall A ge sp st x st' st'' f, + sem ge sp st f st' -> + @step_instr A ge sp st' x st'' -> + sem ge sp st (update f x) st''. +Proof. + intros. + destruct x; inv H0; + eauto using sem_update_RBnop, sem_update2_Op, sem_update2_load, sem_update2_store. +Qed. + +Lemma abstr_sem_val_mem : + forall A B ge tge st tst sp a, + ge_preserved ge tge -> + forall v m, + (@sem_mem A ge sp st a m /\ match_states st tst -> @sem_mem B tge sp tst a m) /\ + (@sem_value A ge sp st a v /\ match_states st tst -> @sem_value B tge sp tst a v). +Proof. + intros * H. + apply expression_ind2 with + + (P := fun (e1: expression) => + forall v m, + (@sem_mem A ge sp st e1 m /\ match_states st tst -> @sem_mem B tge sp tst e1 m) /\ + (@sem_value A ge sp st e1 v /\ match_states st tst -> @sem_value B tge sp tst e1 v)) + + (P0 := fun (e1: expression_list) => + forall lv, @sem_val_list A ge sp st e1 lv /\ match_states st tst -> @sem_val_list B tge sp tst e1 lv); + simplify; intros; simplify. + { inv H1. inv H2. constructor. } + { inv H2. inv H1. rewrite H0. constructor. } + { inv H3. } + { inv H3. inv H4. econstructor. apply H1; auto. simplify. eauto. constructor. auto. auto. + apply H0; simplify; eauto. constructor; eauto. + unfold ge_preserved in *. simplify. rewrite <- H2. auto. + } + { inv H3. } + { inv H3. inv H4. econstructor. apply H1; eauto; simplify; eauto. constructor; eauto. + apply H0; simplify; eauto. constructor; eauto. + inv H. rewrite <- H4. eauto. + auto. + } + { inv H4. inv H5. econstructor. apply H0; eauto. simplify; eauto. constructor; eauto. + apply H2; eauto. simplify; eauto. constructor; eauto. + apply H1; eauto. simplify; eauto. constructor; eauto. + inv H. rewrite <- H5. eauto. auto. + } + { inv H4. } + { inv H1. constructor. } + { inv H3. constructor; auto. apply H0; eauto. apply Mem.empty. } +Qed. + +Lemma abstr_sem_value : + forall a A B ge tge sp st tst v, + @sem_value A ge sp st a v -> + ge_preserved ge tge -> + match_states st tst -> + @sem_value B tge sp tst a v. +Proof. intros; eapply abstr_sem_val_mem; eauto; apply Mem.empty. Qed. + +Lemma abstr_sem_mem : + forall a A B ge tge sp st tst v, + @sem_mem A ge sp st a v -> + ge_preserved ge tge -> + match_states st tst -> + @sem_mem B tge sp tst a v. +Proof. intros; eapply abstr_sem_val_mem; eauto. Qed. + +Lemma abstr_sem_regset : + forall a a' A B ge tge sp st tst rs, + @sem_regset A ge sp st a rs -> + ge_preserved ge tge -> + (forall x, a # x = a' # x) -> + match_states st tst -> + exists rs', @sem_regset B tge sp tst a' rs' /\ (forall x, rs !! x = rs' !! x). +Proof. + inversion 1; intros. + inv H7. + econstructor. simplify. econstructor. intros. + eapply abstr_sem_value; eauto. rewrite <- H6. + eapply H0. constructor; eauto. + auto. +Qed. + +Lemma abstr_sem : + forall a a' A B ge tge sp st tst st', + @sem A ge sp st a st' -> + ge_preserved ge tge -> + (forall x, a # x = a' # x) -> + match_states st tst -> + exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. +Proof. + inversion 1; subst; intros. + inversion H4; subst. + exploit abstr_sem_regset; eauto; inv_simp. + do 3 econstructor; eauto. + rewrite <- H3. + eapply abstr_sem_mem; eauto. +Qed. + +Lemma abstract_execution_correct': + forall A B ge tge sp st' a a' st tst, + @sem A ge sp st a st' -> + ge_preserved ge tge -> + check a a' = true -> + match_states st tst -> + exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. +Proof. + intros; + pose proof (check_correct a a' H1); + eapply abstr_sem; eauto. +Qed. + +Lemma states_match : + forall st1 st2 st3 st4, + match_states st1 st2 -> + match_states st2 st3 -> + match_states st3 st4 -> + match_states st1 st4. +Proof. + intros * H1 H2 H3; destruct st1; destruct st2; destruct st3; destruct st4. + inv H1. inv H2. inv H3; constructor. + unfold regs_lessdef in *. intros. + repeat match goal with + | H: forall _, _, r : positive |- _ => specialize (H r) + end. + congruence. + auto. +Qed. + +Lemma step_instr_block_same : + forall ge sp st st', + step_instr_block ge sp st nil st' -> + st = st'. +Proof. inversion 1; auto. Qed. + +Lemma step_instr_seq_same : + forall ge sp st st', + step_instr_seq ge sp st nil st' -> + st = st'. +Proof. inversion 1; auto. Qed. + +Lemma sem_update' : + forall A ge sp st a x st', + sem ge sp st (update (abstract_sequence empty a) x) st' -> + exists st'', + @step_instr A ge sp st'' x st' /\ + sem ge sp st (abstract_sequence empty a) st''. +Proof. + Admitted. + +Lemma rtlpar_trans_correct : + forall bb ge sp sem_st' sem_st st, + sem ge sp sem_st (abstract_sequence empty (concat (concat bb))) sem_st' -> + match_states sem_st st -> + exists st', RTLPar.step_instr_block ge sp st bb st' + /\ match_states sem_st' st'. +Proof. + induction bb using rev_ind. + { repeat econstructor. eapply abstract_interp_empty3 in H. + inv H. inv H0. constructor; congruence. } + { simplify. inv H0. repeat rewrite concat_app in H. simplify. + rewrite app_nil_r in H. + exploit sem_separate; eauto; inv_simp. + repeat econstructor. admit. admit. + } +Admitted. + +(*Lemma abstract_execution_correct_ld: + forall bb bb' cfi ge tge sp st st' tst, + RTLBlock.step_instr_list ge sp st bb st' -> + ge_preserved ge tge -> + schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi) = true -> + match_states_ld st tst -> + exists tst', RTLPar.step_instr_block tge sp tst bb' tst' + /\ match_states st' tst'. +Proof. + intros.*) +*) + +Lemma match_states_list : + forall A (rs: Regmap.t A) rs', + (forall r, rs !! r = rs' !! r) -> + forall l, rs ## l = rs' ## l. +Proof. induction l; crush. Qed. + +Lemma PTree_matches : + forall A (v: A) res rs rs', + (forall r, rs !! r = rs' !! r) -> + forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x. +Proof. + intros; destruct (Pos.eq_dec x res); subst; + [ repeat rewrite Regmap.gss by auto + | repeat rewrite Regmap.gso by auto ]; auto. +Qed. + +Lemma abstract_interp_empty3 : + forall A ctx st', + @sem A ctx empty st' -> match_states (ctx_is ctx) st'. +Proof. + inversion 1; subst; simplify. destruct ctx. + destruct ctx_is. + constructor; intros. + - inv H0. specialize (H3 x). inv H3. inv H8. reflexivity. + - inv H1. specialize (H3 x). inv H3. inv H8. reflexivity. + - inv H2. inv H8. reflexivity. +Qed. + +Lemma step_instr_matches : + forall A a ge sp st st', + @step_instr A ge sp st a st' -> + forall tst, + match_states st tst -> + exists tst', step_instr ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction 1; simplify; + match goal with H: match_states _ _ |- _ => inv H end; + try solve [repeat econstructor; try erewrite match_states_list; + try apply PTree_matches; eauto; + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end]. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + apply PTree_matches; assumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + econstructor; auto. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + apply PTree_matches; assumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + erewrite <- eval_predf_pr_equiv; eassumption. + econstructor; auto. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + erewrite <- eval_predf_pr_equiv; eassumption. + repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + erewrite <- eval_predf_pr_equiv; eassumption. + match goal with H: eval_pred _ _ _ _ |- _ => inv H end. + repeat econstructor; try erewrite match_states_list; eauto. + match goal with + H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto + end. + - admit. Admitted. + +Lemma step_instr_list_matches : + forall a ge sp st st', + step_instr_list ge sp st a st' -> + forall tst, match_states st tst -> + exists tst', step_instr_list ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction a; intros; inv H; + try (exploit step_instr_matches; eauto; []; simplify; + exploit IHa; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma step_instr_seq_matches : + forall a ge sp st st', + step_instr_seq ge sp st a st' -> + forall tst, match_states st tst -> + exists tst', step_instr_seq ge sp tst a tst' + /\ match_states st' tst'. +Proof. + induction a; intros; inv H; + try (exploit step_instr_list_matches; eauto; []; simplify; + exploit IHa; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma step_instr_block_matches : + forall bb ge sp st st', + step_instr_block ge sp st bb st' -> + forall tst, match_states st tst -> + exists tst', step_instr_block ge sp tst bb tst' + /\ match_states st' tst'. +Proof. + induction bb; intros; inv H; + try (exploit step_instr_seq_matches; eauto; []; simplify; + exploit IHbb; eauto; []; simplify); repeat econstructor; eauto. +Qed. + +Lemma rtlblock_trans_correct' : + forall bb ge sp st x st'', + RTLBlock.step_instr_list ge sp st (bb ++ x :: nil) st'' -> + exists st', RTLBlock.step_instr_list ge sp st bb st' + /\ step_instr ge sp st' x st''. +Proof. + induction bb. + crush. exists st. + split. constructor. inv H. inv H6. auto. + crush. inv H. exploit IHbb. eassumption. simplify. + econstructor. split. + econstructor; eauto. eauto. +Qed. + +Lemma abstract_interp_empty A st : @sem A st empty (ctx_is st). +Proof. destruct st, ctx_is. simpl. repeat econstructor. Qed. + +Lemma abstract_seq : + forall l f i, + abstract_sequence f (l ++ i :: nil) = update (abstract_sequence f l) i. +Proof. induction l; crush. Qed. + +Lemma abstract_sequence_update : + forall l r f, + check_dest_l l r = false -> + (abstract_sequence f l) # (Reg r) = f # (Reg r). +Proof. + induction l using rev_ind; crush. + rewrite abstract_seq. rewrite check_dest_update. apply IHl. + apply check_list_l_false in H. tauto. + apply check_list_l_false in H. tauto. +Qed. + +(*Lemma sem_separate : + forall A ctx b a st', + sem ctx (abstract_sequence empty (a ++ b)) st' -> + exists st'', + @sem A ctx (abstract_sequence empty a) st'' + /\ @sem A (mk_ctx st'' (ctx_sp ctx) (ctx_ge ctx)) (abstract_sequence empty b) st'. +Proof. + induction b using rev_ind; simplify. + { econstructor. simplify. rewrite app_nil_r in H. eauto. apply abstract_interp_empty. } + { simplify. rewrite app_assoc in H. rewrite abstract_seq in H. + exploit sem_update'; eauto; simplify. + exploit IHb; eauto; inv_simp. + econstructor; split; eauto. + rewrite abstract_seq. + eapply sem_update2; eauto. + } +Qed.*) + +Lemma sem_update_RBnop : + forall A ctx f st', + @sem A ctx f st' -> sem ctx (update f RBnop) st'. +Proof. auto. Qed. + +Lemma sem_update_Op : + forall A ge sp ist f st' r l o0 o m rs v ps, + @sem A (mk_ctx ist sp ge) f st' -> + eval_predf ps o = true -> + Op.eval_operation ge sp o0 (rs ## l) m = Some v -> + match_states st' (mk_instr_state rs ps m) -> + exists tst, + sem (mk_ctx ist sp ge) (update f (RBop (Some o) o0 l r)) tst + /\ match_states (mk_instr_state (Regmap.set r v rs) ps m) tst. +Proof. + intros. inv H1. inv H. inv H1. inv H3. simplify. + econstructor. simplify. + { constructor; try constructor; intros; try solve [rewrite genmap1; now eauto]. + destruct (Pos.eq_dec x r); subst. + { rewrite map2. specialize (H1 r). inv H1. +(*} + } + destruct st. + econstructor. simplify. + { constructor. + { constructor. intros. destruct (Pos.eq_dec x r); subst. + { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. + { inv H9. eapply gen_list_base; eauto. } + { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } + { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } + { inv H. rewrite genmap1; crush. eauto. } } + { constructor; eauto. intros. + destruct (Pos.eq_dec r x); + subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } +Qed.*) Admitted. + +Lemma sem_update : + forall A ge sp st x st' st'' st''' f, + sem (mk_ctx st sp ge) f st' -> + 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. + 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', + RTLBlock.step_instr_list ge sp st bb st' -> + forall tst, + match_states st tst -> + exists tst', sem (mk_ctx tst sp ge) (abstract_sequence empty bb) tst' + /\ match_states st' tst'. +Proof. + induction bb using rev_ind; simplify. + { econstructor. simplify. apply abstract_interp_empty. + inv H. auto. } + { apply rtlblock_trans_correct' in H. simplify. + rewrite abstract_seq. + exploit IHbb; try eassumption; []; simplify. + exploit sem_update. apply H1. symmetry; eassumption. + eauto. simplify. econstructor. split. apply H3. + auto. } +Qed. + +Lemma abstract_execution_correct: + forall bb bb' cfi cfi' ge tge sp st st' tst, + RTLBlock.step_instr_list ge sp st bb st' -> + ge_preserved ge tge -> + schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi') = true -> + match_states st tst -> + exists tst', RTLPar.step_instr_block tge sp tst bb' tst' + /\ match_states st' tst'. +Proof. + intros. + unfold schedule_oracle in *. simplify. unfold empty_trees in H4. + exploit rtlblock_trans_correct; try eassumption; []; simplify. +(*) exploit abstract_execution_correct'; + try solve [eassumption | apply state_lessdef_match_sem; eassumption]. + apply match_states_commut. eauto. inv_simp. + exploit rtlpar_trans_correct; try eassumption; []; inv_simp. + exploit step_instr_block_matches; eauto. apply match_states_commut; eauto. inv_simp. + repeat match goal with | H: match_states _ _ |- _ => inv H end. + do 2 econstructor; eauto. + econstructor; congruence. +Qed.*)Admitted. + +Definition match_prog (prog : RTLBlock.program) (tprog : RTLPar.program) := + match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq prog tprog. + +Inductive match_stackframes: RTLBlock.stackframe -> RTLPar.stackframe -> Prop := +| match_stackframe: + forall f tf res sp pc rs rs' ps ps', + transl_function f = OK tf -> + (forall x, rs !! x = rs' !! x) -> + (forall x, ps !! x = ps' !! x) -> + match_stackframes (Stackframe res f sp pc rs ps) + (Stackframe res tf sp pc rs' ps'). + +Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop := +| match_state: + forall sf f sp pc rs rs' m sf' tf ps ps' + (TRANSL: transl_function f = OK tf) + (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_returnstate: + forall stack stack' v m + (STACKS: list_forall2 match_stackframes stack stack'), + match_states (Returnstate stack v m) + (Returnstate stack' v m) +| match_callstate: + forall stack stack' f tf args m + (TRANSL: transl_fundef f = OK tf) + (STACKS: list_forall2 match_stackframes stack stack'), + match_states (Callstate stack f args m) + (Callstate stack' tf args m). + +Section CORRECTNESS. + + Context (prog: RTLBlock.program) (tprog : RTLPar.program). + Context (TRANSL: match_prog prog tprog). + + Let ge : RTLBlock.genv := Globalenvs.Genv.globalenv prog. + Let tge : RTLPar.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. + Hint Resolve symbols_preserved : rtlgp. + + Lemma function_ptr_translated: + forall (b: Values.block) (f: RTLBlock.fundef), + Genv.find_funct_ptr ge b = Some f -> + exists tf, + Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = Errors.OK tf. + Proof using TRANSL. + intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. + intros (cu & tf & P & Q & R); exists tf; auto. + Qed. + + Lemma functions_translated: + forall (v: Values.val) (f: RTLBlock.fundef), + Genv.find_funct ge v = Some f -> + exists tf, + Genv.find_funct tge v = Some tf /\ transl_fundef f = Errors.OK tf. + Proof using TRANSL. + intros. exploit (Genv.find_funct_match TRANSL); eauto. + intros (cu & tf & P & Q & R); exists tf; auto. + Qed. + + Lemma senv_preserved: + Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). + Proof (Genv.senv_transf_partial TRANSL). + Hint Resolve senv_preserved : rtlgp. + + Lemma sig_transl_function: + forall (f: RTLBlock.fundef) (tf: RTLPar.fundef), + transl_fundef f = OK tf -> + funsig tf = funsig f. + Proof using . + unfold transl_fundef, transf_partial_fundef, transl_function; intros; + repeat destruct_match; crush; + match goal with H: OK _ = OK _ |- _ => inv H end; auto. + Qed. + Hint Resolve sig_transl_function : rtlgp. + + Hint Resolve Val.lessdef_same : rtlgp. + Hint Resolve regs_lessdef_regs : rtlgp. + + Lemma find_function_translated: + forall ros rs rs' f, + (forall x, rs !! x = rs' !! x) -> + find_function ge ros rs = Some f -> + exists tf, find_function tge ros rs' = Some tf + /\ transl_fundef f = OK tf. + Proof using TRANSL. + Ltac ffts := match goal with + | [ H: forall _, Val.lessdef _ _, r: Registers.reg |- _ ] => + specialize (H r); inv H + | [ H: Vundef = ?r, H1: Genv.find_funct _ ?r = Some _ |- _ ] => + rewrite <- H in H1 + | [ H: Genv.find_funct _ Vundef = Some _ |- _] => solve [inv H] + | _ => solve [exploit functions_translated; eauto] + end. + destruct ros; simplify; try rewrite <- H; + [| rewrite symbols_preserved; destruct_match; + try (apply function_ptr_translated); crush ]; + intros; + repeat ffts. + Qed. + + Lemma schedule_oracle_nil: + forall bb cfi, + schedule_oracle {| bb_body := nil; bb_exit := cfi |} bb = true -> + bb_body bb = nil /\ bb_exit bb = cfi. + Proof using . + unfold schedule_oracle, check_control_flow_instr. + simplify; repeat destruct_match; crush. + Qed. + + Lemma schedule_oracle_nil2: + forall cfi, + schedule_oracle {| bb_body := nil; bb_exit := cfi |} + {| bb_body := nil; bb_exit := cfi |} = true. + Proof using . + unfold schedule_oracle, check_control_flow_instr. + simplify; repeat destruct_match; crush. + Qed. + + Lemma eval_op_eq: + forall (sp0 : Values.val) (op : Op.operation) (vl : list Values.val) m, + Op.eval_operation ge sp0 op vl m = Op.eval_operation tge sp0 op vl m. + Proof using TRANSL. + intros. + destruct op; auto; unfold Op.eval_operation, Genv.symbol_address, Op.eval_addressing32; + [| destruct a; unfold Genv.symbol_address ]; + try rewrite symbols_preserved; auto. + Qed. + Hint Resolve eval_op_eq : rtlgp. + + Lemma eval_addressing_eq: + forall sp addr vl, + Op.eval_addressing ge sp addr vl = Op.eval_addressing tge sp addr vl. + Proof using TRANSL. + intros. + destruct addr; + unfold Op.eval_addressing, Op.eval_addressing32; + unfold Genv.symbol_address; + try rewrite symbols_preserved; auto. + Qed. + Hint Resolve eval_addressing_eq : rtlgp. + + Lemma ge_preserved_lem: + ge_preserved ge tge. + Proof using TRANSL. + unfold ge_preserved. + eauto with rtlgp. + Qed. + Hint Resolve ge_preserved_lem : rtlgp. + + Lemma lessdef_regmap_optget: + forall or rs rs', + regs_lessdef rs rs' -> + Val.lessdef (regmap_optget or Vundef rs) (regmap_optget or Vundef rs'). + Proof using. destruct or; crush. Qed. + Hint Resolve lessdef_regmap_optget : rtlgp. + + Lemma regmap_equiv_lessdef: + forall rs rs', + (forall x, rs !! x = rs' !! x) -> + regs_lessdef rs rs'. + Proof using. + intros; unfold regs_lessdef; intros. + rewrite H. apply Val.lessdef_refl. + Qed. + Hint Resolve regmap_equiv_lessdef : rtlgp. + + Lemma int_lessdef: + forall rs rs', + regs_lessdef rs rs' -> + (forall arg v, + rs !! arg = Vint v -> + rs' !! arg = Vint v). + Proof using. intros ? ? H; intros; specialize (H arg); inv H; crush. Qed. + Hint Resolve int_lessdef : rtlgp. + + Ltac semantics_simpl := + match goal with + | [ H: match_states _ _ |- _ ] => + let H2 := fresh "H" in + learn H as H2; inv H2 + | [ H: transl_function ?f = OK _ |- _ ] => + let H2 := fresh "TRANSL" in + learn H as H2; + unfold transl_function in H2; + destruct (check_scheduled_trees + (@fn_code RTLBlock.bb f) + (@fn_code RTLPar.bb (schedule f))) eqn:?; + [| discriminate ]; inv H2 + | [ H: context[check_scheduled_trees] |- _ ] => + let H2 := fresh "CHECK" in + learn H as H2; + eapply check_scheduled_trees_correct in H2; [| solve [eauto] ] + | [ H: schedule_oracle {| bb_body := nil; bb_exit := _ |} ?bb = true |- _ ] => + let H2 := fresh "SCHED" in + learn H as H2; + apply schedule_oracle_nil in H2 + | [ H: find_function _ _ _ = Some _, H2: forall x, ?rs !! x = ?rs' !! x |- _ ] => + learn H; exploit find_function_translated; try apply H2; eauto; inversion 1 + | [ H: Mem.free ?m _ _ _ = Some ?m', H2: Mem.extends ?m ?m'' |- _ ] => + learn H; exploit Mem.free_parallel_extends; eauto; intros + | [ H: Events.eval_builtin_args _ _ _ _ _ _, H2: regs_lessdef ?rs ?rs' |- _ ] => + let H3 := fresh "H" in + learn H; exploit Events.eval_builtin_args_lessdef; [apply H2 | | |]; + eauto with rtlgp; intro H3; learn H3 + | [ H: Events.external_call _ _ _ _ _ _ _ |- _ ] => + let H2 := fresh "H" in + learn H; exploit Events.external_call_mem_extends; + eauto; intro H2; learn H2 + | [ H: exists _, _ |- _ ] => inv H + | _ => progress simplify + end. + + Hint Resolve Events.eval_builtin_args_preserved : rtlgp. + Hint Resolve Events.external_call_symbols_preserved : rtlgp. + Hint Resolve set_res_lessdef : rtlgp. + Hint Resolve set_reg_lessdef : rtlgp. + Hint Resolve Op.eval_condition_lessdef : rtlgp. + + Hint Constructors Events.eval_builtin_arg: barg. + + Lemma eval_builtin_arg_eq: + forall A ge a v1 m1 e1 e2 sp, + (forall x, e1 x = e2 x) -> + @Events.eval_builtin_arg A ge e1 sp m1 a v1 -> + Events.eval_builtin_arg ge e2 sp m1 a v1. +Proof. induction 2; try rewrite H; eauto with barg. Qed. + + Lemma eval_builtin_args_eq: + forall A ge e1 sp m1 e2 al vl1, + (forall x, e1 x = e2 x) -> + @Events.eval_builtin_args A ge e1 sp m1 al vl1 -> + Events.eval_builtin_args ge e2 sp m1 al vl1. + Proof. + induction 2. + - econstructor; split. + - exploit eval_builtin_arg_eq; eauto. intros. + destruct IHlist_forall2 as [| y]. constructor; eauto. + constructor. constructor; auto. + constructor; eauto. + Qed. + + Lemma step_cf_instr_correct: + forall cfi t s s', + step_cf_instr ge s cfi t s' -> + forall r, + match_states s r -> + exists r', step_cf_instr tge r cfi t r' /\ match_states s' r'. + Proof using TRANSL. + induction 1; repeat semantics_simpl; + try solve [repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp)]. + { do 3 (try erewrite match_states_list by eauto; econstructor; eauto with rtlgp). + eapply eval_builtin_args_eq. eapply REG. + eapply Events.eval_builtin_args_preserved. eapply symbols_preserved. + eauto. + intros. + unfold regmap_setres. destruct res. + destruct (Pos.eq_dec x0 x); subst. + repeat rewrite Regmap.gss; auto. + repeat rewrite Regmap.gso; auto. + eapply REG. eapply REG. + } + { repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). + unfold regmap_optget. destruct or. rewrite REG. constructor; eauto. + constructor; eauto. + } + { exploit IHstep_cf_instr; eauto. simplify. + repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). + erewrite eval_predf_pr_equiv; eauto. + } + Qed. + + Theorem transl_step_correct : + forall (S1 : RTLBlock.state) t S2, + RTLBlock.step ge S1 t S2 -> + forall (R1 : RTLPar.state), + match_states S1 R1 -> + exists R2, Smallstep.plus RTLPar.step tge R1 t R2 /\ match_states S2 R2. + Proof. + + induction 1; repeat semantics_simpl. + + { destruct bb; destruct x. + assert (bb_exit = bb_exit0). + { unfold schedule_oracle in *. simplify. + unfold check_control_flow_instr in *. + destruct_match; crush. + } + subst. + + exploit abstract_execution_correct; try eassumption. eapply ge_preserved_lem. + econstructor; eauto. + simplify. destruct x. inv H7. + + exploit step_cf_instr_correct; try eassumption. econstructor; eauto. + simplify. + + econstructor. econstructor. eapply Smallstep.plus_one. econstructor. + 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. + 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. } + { inv STACKS. inv H2. repeat econstructor; eauto. + intros. apply PTree_matches; eauto. } + Qed. + + Lemma transl_initial_states: + forall S, + RTLBlock.initial_state prog S -> + exists R, RTLPar.initial_state tprog R /\ match_states S R. + Proof. + induction 1. + exploit function_ptr_translated; eauto. intros [tf [A B]]. + econstructor; split. + econstructor. apply (Genv.init_mem_transf_partial TRANSL); eauto. + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto. + symmetry; eapply match_program_main; eauto. + eexact A. + rewrite <- H2. apply sig_transl_function; auto. + constructor. auto. constructor. + Qed. + + Lemma transl_final_states: + forall S R r, + match_states S R -> RTLBlock.final_state S r -> RTLPar.final_state R r. + Proof. + intros. inv H0. inv H. inv STACKS. constructor. + Qed. + + Theorem transf_program_correct: + Smallstep.forward_simulation (RTLBlock.semantics prog) (RTLPar.semantics tprog). + Proof. + eapply Smallstep.forward_simulation_plus. + apply senv_preserved. + eexact transl_initial_states. + eexact transl_final_states. + exact transl_step_correct. + Qed. + +End CORRECTNESS. +#+end_src + +* License + +#+name: license +#+begin_src coq :tangle no +(* + * 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 . + *) +#+end_src diff --git a/lit/basic-block-generation.org b/lit/basic-block-generation.org deleted file mode 100644 index b835e61..0000000 --- a/lit/basic-block-generation.org +++ /dev/null @@ -1,378 +0,0 @@ -#+title: Basic Block Generation -#+author: Yann Herklotz -#+email: yann [at] yannherklotz [dot] com - -* RTLBlockgen -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgen.v -:END: - -Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]]. - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblockgen-imports -#+begin_src coq -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. -#+end_src - -#+name: rtlblockgen-equalities-insert -#+begin_src coq :comments no -<> -#+end_src - -#+name: rtlblockgen-main -#+begin_src coq -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 - | 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. -#+end_src - -** Equalities - -#+name: rtlblockgen-equalities -#+begin_src coq :tangle no -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. -#+end_src - -* RTLBlockgenproof -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -** Imports - -#+name: rtlblockgenproof-imports -#+begin_src coq -Require compcert.backend.RTL. -Require Import compcert.common.AST. -Require Import compcert.lib.Maps. - -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLBlockgen. -#+end_src - -** Match states - -The ~match_states~ predicate describes which states are equivalent between the two languages, in this -case ~RTL~ and ~RTLBlock~. - -#+name: rtlblockgenproof-match-states -#+begin_src coq -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). -#+end_src - -** Correctness - -#+name: rtlblockgenproof-correctness -#+begin_src coq -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. -#+end_src - -* License - -#+name: license -#+begin_src coq :tangle no -(* - * 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 . - *) -#+end_src diff --git a/lit/scheduler-languages.org b/lit/scheduler-languages.org deleted file mode 100644 index cf03b3a..0000000 --- a/lit/scheduler-languages.org +++ /dev/null @@ -1,689 +0,0 @@ -#+title: Scheduler Languages -#+author: Yann Herklotz -#+email: yann [at] yannherklotz [dot] com - -* RTLBlockInstr -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockInstr.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblockinstr-imports -#+begin_src coq -Require Import Coq.micromega.Lia. - -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.Values. -Require Import compcert.lib.Integers. -Require Import compcert.lib.Maps. -Require Import compcert.verilog.Op. - -Require Import Predicate. -Require Import Vericertlib. -#+end_src - -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]). - -#+name: rtlblockinstr-instr-def -#+begin_src coq -Definition node := positive. - -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. -#+end_src - -** Control-Flow Instruction Definition - -These are the instructions that count as control-flow, and will be placed at the end of the basic -blocks. - -#+name: rtlblockinstr-cf-instr-def -#+begin_src coq -Inductive cf_instr : Type := -| RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr -| RBtailcall : signature -> reg + ident -> list reg -> cf_instr -| RBbuiltin : external_function -> list (builtin_arg reg) -> - builtin_res reg -> node -> cf_instr -| RBcond : condition -> list reg -> node -> node -> cf_instr -| RBjumptable : reg -> list node -> cf_instr -| RBreturn : option reg -> cf_instr -| RBgoto : node -> cf_instr -| RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. -#+end_src - -** Helper functions - -#+name: rtlblockinstr-helpers -#+begin_src coq -Fixpoint successors_instr (i : cf_instr) : list node := - match i with - | RBcall sig ros args res s => s :: nil - | RBtailcall sig ros args => nil - | RBbuiltin ef args res s => s :: nil - | RBcond cond args ifso ifnot => ifso :: ifnot :: nil - | RBjumptable arg tbl => tbl - | RBreturn optarg => nil - | RBgoto n => n :: nil - | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil) - end. - -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) - | RBload p chunk addr args dst => - fold_left Pos.max args (Pos.max dst m) - | RBstore p chunk addr args src => - fold_left Pos.max args (Pos.max src m) - | RBsetpred p' c args p => - 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)) - | RBcall sig (inr id) args res s => - fold_left Pos.max args (Pos.max res m) - | RBtailcall sig (inl r) args => - fold_left Pos.max args (Pos.max r m) - | RBtailcall sig (inr id) args => - 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) - | RBcond cond args ifso ifnot => fold_left Pos.max args m - | RBjumptable arg tbl => Pos.max arg m - | RBreturn None => m - | RBreturn (Some arg) => Pos.max arg m - | RBgoto n => m - | RBpred_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. -#+end_src - -*** 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. - -#+name: rtlblockinstr-instr-state -#+begin_src coq -Record instr_state := mk_instr_state { - is_rs: regset; - is_ps: predset; - is_mem: mem; -}. -#+end_src - -** Top-Level Type Definitions - -#+name: rtlblockinstr-type-def -#+begin_src coq -Section DEFINITION. - - Context {bblock_body: Type}. - - 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_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. - -End DEFINITION. -#+end_src - -** Semantics - -#+name: rtlblockinstr-semantics -#+begin_src coq -Section RELSEM. - - Context {bblock_body : Type}. - - Definition genv := Genv.t (@fundef bblock_body) unit. - - 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_RBnop: - 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 - | 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 - | 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 - | 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. - - 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, - 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') - 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) (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') - | 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) - | 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) - | 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) - 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) - | 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'. - -End RELSEM. -#+end_src - -* RTLBlock -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlock.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblock-main -#+begin_src coq -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.RTLBlockInstr. - -Definition bb := list instr. - -Definition bblock := @bblock bb. -Definition code := @code bb. -Definition function := @function bb. -Definition fundef := @fundef bb. -Definition program := @program bb. -Definition funsig := @funsig bb. -Definition stackframe := @stackframe 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). - - 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_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). -#+end_src - -* RTLPar -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPar.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpar-main -#+begin_src coq -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.RTLBlockInstr. - -Definition bb := list (list (list instr)). - -Definition bblock := @bblock bb. -Definition code := @code bb. -Definition function := @function bb. -Definition fundef := @fundef bb. -Definition program := @program bb. -Definition funsig := @funsig bb. -Definition stackframe := @stackframe bb. -Definition state := @state bb. -Definition genv := @genv bb. - -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_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 -> bb -> 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: 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 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') - | 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. -#+end_src - -* License - -#+name: license -#+begin_src coq :tangle no -(* - * 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 . - *) -#+end_src diff --git a/lit/scheduler.org b/lit/scheduler.org deleted file mode 100644 index 018633c..0000000 --- a/lit/scheduler.org +++ /dev/null @@ -1,2295 +0,0 @@ -#+title: Basic Block Generation -#+author: Yann Herklotz -#+email: yann [at] yannherklotz [dot] com - -* Scheduler -:PROPERTIES: -:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Schedule.ml -:END: - -#+begin_src ocaml :comments no :padline no :exports none -<> -#+end_src - -#+name: scheduler-main -#+begin_src ocaml -open Printf -open Clflags -open Camlcoq -open Datatypes -open Coqlib -open Maps -open AST -open Kildall -open Op -open RTLBlockInstr -open Predicate -open RTLBlock -open HTL -open Verilog -open HTLgen -open HTLMonad -open HTLMonadExtra - -module SS = Set.Make(P) - -type svtype = - | BBType of int - | VarType of int * int - -type sv = { - sv_type: svtype; - sv_num: int; -} - -let print_sv v = - match v with - | { sv_type = BBType bbi; sv_num = n } -> sprintf "bb%d_%d" bbi n - | { sv_type = VarType (bbi, i); sv_num = n } -> sprintf "var%dn%d_%d" bbi i n - -module G = Graph.Persistent.Digraph.ConcreteLabeled(struct - type t = sv - let compare = compare - let equal = (=) - let hash = Hashtbl.hash -end)(struct - type t = int - let compare = compare - let hash = Hashtbl.hash - let equal = (=) - let default = 0 -end) - -module GDot = Graph.Graphviz.Dot(struct - let graph_attributes _ = [] - let default_vertex_attributes _ = [] - let vertex_name = print_sv - let vertex_attributes _ = [] - let get_subgraph _ = None - let default_edge_attributes _ = [] - let edge_attributes _ = [] - - include G - end) - -module DFG = Graph.Persistent.Digraph.ConcreteLabeled(struct - type t = int * instr - let compare = compare - let equal = (=) - let hash = Hashtbl.hash -end)(struct - type t = int - let compare = compare - let hash = Hashtbl.hash - let equal = (=) - let default = 0 -end) - -module DFGSimp = Graph.Persistent.Graph.Concrete(struct - type t = int * instr - let compare = compare - let equal = (=) - let hash = Hashtbl.hash - end) - -let convert dfg = - DFG.fold_vertex (fun v g -> DFGSimp.add_vertex g v) dfg DFGSimp.empty - |> DFG.fold_edges (fun v1 v2 g -> DFGSimp.add_edge (DFGSimp.add_edge g v1 v2) v2 v1) dfg - -let reg r = sprintf "r%d" (P.to_int r) -let print_pred r = sprintf "p%d" (P.to_int r) - -let print_instr = function - | RBnop -> "" - | RBload (_, _, _, _, r) -> sprintf "load(%s)" (reg r) - | RBstore (_, _, _, _, r) -> sprintf "store(%s)" (reg r) - | RBsetpred (_, _, _, p) -> sprintf "setpred(%s)" (print_pred p) - | RBop (_, op, args, d) -> - (match op, args with - | Omove, _ -> "mov" - | Ointconst n, _ -> sprintf "%s=%ld" (reg d) (camlint_of_coqint n) - | Olongconst n, _ -> sprintf "%s=%LdL" (reg d) (camlint64_of_coqint n) - | Ofloatconst n, _ -> sprintf "%s=%.15F" (reg d) (camlfloat_of_coqfloat n) - | Osingleconst n, _ -> sprintf "%s=%.15Ff" (reg d) (camlfloat_of_coqfloat32 n) - | Oindirectsymbol id, _ -> sprintf "%s=&%s" (reg d) (extern_atom id) - | Ocast8signed, [r1] -> sprintf "%s=int8signed(%s)" (reg d) (reg r1) - | Ocast8unsigned, [r1] -> sprintf "%s=int8unsigned(%s)" (reg d) (reg r1) - | Ocast16signed, [r1] -> sprintf "%s=int16signed(%s)" (reg d) (reg r1) - | Ocast16unsigned, [r1] -> sprintf "%s=int16unsigned(%s)" (reg d) (reg r1) - | Oneg, [r1] -> sprintf "%s=-%s" (reg d) (reg r1) - | Osub, [r1;r2] -> sprintf "%s=%s-%s" (reg d) (reg r1) (reg r2) - | Omul, [r1;r2] -> sprintf "%s=%s*%s" (reg d) (reg r1) (reg r2) - | Omulimm n, [r1] -> sprintf "%s=%s*%ld" (reg d) (reg r1) (camlint_of_coqint n) - | Omulhs, [r1;r2] -> sprintf "%s=mulhs(%s,%s)" (reg d) (reg r1) (reg r2) - | Omulhu, [r1;r2] -> sprintf "%s=mulhu(%s,%s)" (reg d) (reg r1) (reg r2) - | Odiv, [r1;r2] -> sprintf "%s=%s /s %s" (reg d) (reg r1) (reg r2) - | Odivu, [r1;r2] -> sprintf "%s=%s /u %s" (reg d) (reg r1) (reg r2) - | Omod, [r1;r2] -> sprintf "%s=%s %%s %s" (reg d) (reg r1) (reg r2) - | Omodu, [r1;r2] -> sprintf "%s=%s %%u %s" (reg d) (reg r1) (reg r2) - | Oand, [r1;r2] -> sprintf "%s=%s & %s" (reg d) (reg r1) (reg r2) - | Oandimm n, [r1] -> sprintf "%s=%s & %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oor, [r1;r2] -> sprintf "%s=%s | %s" (reg d) (reg r1) (reg r2) - | Oorimm n, [r1] -> sprintf "%s=%s | %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oxor, [r1;r2] -> sprintf "%s=%s ^ %s" (reg d) (reg r1) (reg r2) - | Oxorimm n, [r1] -> sprintf "%s=%s ^ %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Onot, [r1] -> sprintf "%s=not(%s)" (reg d) (reg r1) - | Oshl, [r1;r2] -> sprintf "%s=%s << %s" (reg d) (reg r1) (reg r2) - | Oshlimm n, [r1] -> sprintf "%s=%s << %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshr, [r1;r2] -> sprintf "%s=%s >>s %s" (reg d) (reg r1) (reg r2) - | Oshrimm n, [r1] -> sprintf "%s=%s >>s %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrximm n, [r1] -> sprintf "%s=%s >>x %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshru, [r1;r2] -> sprintf "%s=%s >>u %s" (reg d) (reg r1) (reg r2) - | Oshruimm n, [r1] -> sprintf "%s=%s >>u %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Ororimm n, [r1] -> sprintf "%s=%s ror %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshldimm n, [r1;r2] -> sprintf "%s=(%s, %s) << %ld" (reg d) (reg r1) (reg r2) (camlint_of_coqint n) - | Olea addr, args -> sprintf "%s=addr" (reg d) - | Omakelong, [r1;r2] -> sprintf "%s=makelong(%s,%s)" (reg d) (reg r1) (reg r2) - | Olowlong, [r1] -> sprintf "%s=lowlong(%s)" (reg d) (reg r1) - | Ohighlong, [r1] -> sprintf "%s=highlong(%s)" (reg d) (reg r1) - | Ocast32signed, [r1] -> sprintf "%s=long32signed(%s)" (reg d) (reg r1) - | Ocast32unsigned, [r1] -> sprintf "%s=long32unsigned(%s)" (reg d) (reg r1) - | Onegl, [r1] -> sprintf "%s=-l %s" (reg d) (reg r1) - | Osubl, [r1;r2] -> sprintf "%s=%s -l %s" (reg d) (reg r1) (reg r2) - | Omull, [r1;r2] -> sprintf "%s=%s *l %s" (reg d) (reg r1) (reg r2) - | Omullimm n, [r1] -> sprintf "%s=%s *l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Omullhs, [r1;r2] -> sprintf "%s=mullhs(%s,%s)" (reg d) (reg r1) (reg r2) - | Omullhu, [r1;r2] -> sprintf "%s=mullhu(%s,%s)" (reg d) (reg r1) (reg r2) - | Odivl, [r1;r2] -> sprintf "%s=%s /ls %s" (reg d) (reg r1) (reg r2) - | Odivlu, [r1;r2] -> sprintf "%s=%s /lu %s" (reg d) (reg r1) (reg r2) - | Omodl, [r1;r2] -> sprintf "%s=%s %%ls %s" (reg d) (reg r1) (reg r2) - | Omodlu, [r1;r2] -> sprintf "%s=%s %%lu %s" (reg d) (reg r1) (reg r2) - | Oandl, [r1;r2] -> sprintf "%s=%s &l %s" (reg d) (reg r1) (reg r2) - | Oandlimm n, [r1] -> sprintf "%s=%s &l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Oorl, [r1;r2] -> sprintf "%s=%s |l %s" (reg d) (reg r1) (reg r2) - | Oorlimm n, [r1] -> sprintf "%s=%s |l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Oxorl, [r1;r2] -> sprintf "%s=%s ^l %s" (reg d) (reg r1) (reg r2) - | Oxorlimm n, [r1] -> sprintf "%s=%s ^l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Onotl, [r1] -> sprintf "%s=notl(%s)" (reg d) (reg r1) - | Oshll, [r1;r2] -> sprintf "%s=%s < sprintf "%s=%s < sprintf "%s=%s >>ls %s" (reg d) (reg r1) (reg r2) - | Oshrlimm n, [r1] -> sprintf "%s=%s >>ls %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrxlimm n, [r1] -> sprintf "%s=%s >>lx %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrlu, [r1;r2] -> sprintf "%s=%s >>lu %s" (reg d) (reg r1) (reg r2) - | Oshrluimm n, [r1] -> sprintf "%s=%s >>lu %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Ororlimm n, [r1] -> sprintf "%s=%s rorl %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oleal addr, args -> sprintf "%s=addr" (reg d) - | Onegf, [r1] -> sprintf "%s=negf(%s)" (reg d) (reg r1) - | Oabsf, [r1] -> sprintf "%s=absf(%s)" (reg d) (reg r1) - | Oaddf, [r1;r2] -> sprintf "%s=%s +f %s" (reg d) (reg r1) (reg r2) - | Osubf, [r1;r2] -> sprintf "%s=%s -f %s" (reg d) (reg r1) (reg r2) - | Omulf, [r1;r2] -> sprintf "%s=%s *f %s" (reg d) (reg r1) (reg r2) - | Odivf, [r1;r2] -> sprintf "%s=%s /f %s" (reg d) (reg r1) (reg r2) - | Onegfs, [r1] -> sprintf "%s=negfs(%s)" (reg d) (reg r1) - | Oabsfs, [r1] -> sprintf "%s=absfs(%s)" (reg d) (reg r1) - | Oaddfs, [r1;r2] -> sprintf "%s=%s +fs %s" (reg d) (reg r1) (reg r2) - | Osubfs, [r1;r2] -> sprintf "%s=%s -fs %s" (reg d) (reg r1) (reg r2) - | Omulfs, [r1;r2] -> sprintf "%s=%s *fs %s" (reg d) (reg r1) (reg r2) - | Odivfs, [r1;r2] -> sprintf "%s=%s /fs %s" (reg d) (reg r1) (reg r2) - | Osingleoffloat, [r1] -> sprintf "%s=singleoffloat(%s)" (reg d) (reg r1) - | Ofloatofsingle, [r1] -> sprintf "%s=floatofsingle(%s)" (reg d) (reg r1) - | Ointoffloat, [r1] -> sprintf "%s=intoffloat(%s)" (reg d) (reg r1) - | Ofloatofint, [r1] -> sprintf "%s=floatofint(%s)" (reg d) (reg r1) - | Ointofsingle, [r1] -> sprintf "%s=intofsingle(%s)" (reg d) (reg r1) - | Osingleofint, [r1] -> sprintf "%s=singleofint(%s)" (reg d) (reg r1) - | Olongoffloat, [r1] -> sprintf "%s=longoffloat(%s)" (reg d) (reg r1) - | Ofloatoflong, [r1] -> sprintf "%s=floatoflong(%s)" (reg d) (reg r1) - | Olongofsingle, [r1] -> sprintf "%s=longofsingle(%s)" (reg d) (reg r1) - | Osingleoflong, [r1] -> sprintf "%s=singleoflong(%s)" (reg d) (reg r1) - | Ocmp c, args -> sprintf "%s=cmp" (reg d) - | Osel (c, ty), r1::r2::args -> sprintf "%s=sel" (reg d) - | _, _ -> sprintf "N/a") - -module DFGDot = Graph.Graphviz.Dot(struct - let graph_attributes _ = [] - let default_vertex_attributes _ = [] - let vertex_name = function (i, instr) -> sprintf "\"%d:%s\"" i (print_instr instr) - let vertex_attributes _ = [] - let get_subgraph _ = None - let default_edge_attributes _ = [] - let edge_attributes _ = [] - - include DFG - end) - -module DFGDfs = Graph.Traverse.Dfs(DFG) - -module IMap = Map.Make (struct - type t = int - - let compare = compare -end) - -let gen_vertex instrs i = (i, List.nth instrs i) - -(** The DFG type defines a list of instructions with their data dependencies as [edges], which are - the pairs of integers that represent the index of the instruction in the [nodes]. The edges - always point from left to right. *) - -let print_list f out_chan a = - fprintf out_chan "[ "; - List.iter (fprintf out_chan "%a " f) a; - fprintf out_chan "]" - -let print_tuple out_chan a = - let l, r = a in - fprintf out_chan "(%d,%d)" l r - -(*let print_dfg out_chan dfg = - fprintf out_chan "{ nodes = %a, edges = %a }" - (print_list PrintRTLBlockInstr.print_bblock_body) - dfg.nodes (print_list print_tuple) dfg.edges*) - -let print_dfg = DFGDot.output_graph - -let read_process command = - let buffer_size = 2048 in - let buffer = Buffer.create buffer_size in - let string = Bytes.create buffer_size in - let in_channel = Unix.open_process_in command in - let chars_read = ref 1 in - while !chars_read <> 0 do - chars_read := input in_channel string 0 buffer_size; - Buffer.add_substring buffer (Bytes.to_string string) 0 !chars_read - done; - ignore (Unix.close_process_in in_channel); - Buffer.contents buffer - -let comb_delay = function - | RBnop -> 0 - | RBop (_, op, _, _) -> - (match op with - | Omove -> 0 - | Ointconst _ -> 0 - | Olongconst _ -> 0 - | Ocast8signed -> 0 - | Ocast8unsigned -> 0 - | Ocast16signed -> 0 - | Ocast16unsigned -> 0 - | Oneg -> 0 - | Onot -> 0 - | Oor -> 0 - | Oorimm _ -> 0 - | Oand -> 0 - | Oandimm _ -> 0 - | Oxor -> 0 - | Oxorimm _ -> 0 - | Omul -> 8 - | Omulimm _ -> 8 - | Omulhs -> 8 - | Omulhu -> 8 - | Odiv -> 72 - | Odivu -> 72 - | Omod -> 72 - | Omodu -> 72 - | _ -> 1) - | _ -> 1 - -let pipeline_stages = function - | RBload _ -> 2 - | RBop (_, op, _, _) -> - (match op with - | Odiv -> 32 - | Odivu -> 32 - | Omod -> 32 - | Omodu -> 32 - | _ -> 0) - | _ -> 0 - -(** Add a dependency if it uses a register that was written to previously. *) -let add_dep map i tree dfg curr = - match PTree.get curr tree with - | None -> dfg - | Some ip -> - let ipv = (List.nth map ip) in - 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. - - This function only gathers the RAW constraints, and will therefore only be active for operations - that modify registers, which is this case only affects loads and operations. *) -let accumulate_RAW_deps map dfg curr = - let i, dst_map, graph = dfg in - let acc_dep_instruction rs dst = - ( i + 1, - PTree.set dst i dst_map, - List.fold_left (add_dep map i dst_map) graph rs - ) - in - let acc_dep_instruction_nodst rs = - ( i + 1, - dst_map, - List.fold_left (add_dep map i dst_map) graph rs) - in - 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) - -(** Finds the next write to the [dst] register. This is a small optimisation so that only one - dependency is generated for a data dependency. *) -let rec find_next_dst_write i dst i' curr = - let check_dst dst' curr' = - if dst = dst' then Some (i, i') - else find_next_dst_write i dst (i' + 1) curr' - in - match curr with - | [] -> None - | RBop (_, _, _, dst') :: curr' -> check_dst dst' curr' - | RBload (_, _, _, _, dst') :: curr' -> check_dst dst' curr' - | _ :: curr' -> find_next_dst_write i dst (i' + 1) curr' - -let rec find_all_next_dst_read i dst i' curr = - let check_dst rs curr' = - if List.exists (fun x -> x = dst) rs - then (i, i') :: find_all_next_dst_read i dst (i' + 1) curr' - else find_all_next_dst_read i dst (i' + 1) curr' - in - match curr with - | [] -> [] - | RBop (_, _, rs, _) :: curr' -> check_dst rs curr' - | RBload (_, _, _, rs, _) :: curr' -> check_dst rs curr' - | RBstore (_, _, _, rs, src) :: curr' -> check_dst (src :: rs) curr' - | RBnop :: curr' -> find_all_next_dst_read i dst (i' + 1) curr' - | RBsetpred (_, _, rs, _) :: curr' -> check_dst rs curr' - -let drop i lst = - let rec drop' i' lst' = - match lst' with - | _ :: ls -> if i' = i then ls else drop' (i' + 1) ls - | [] -> [] - in - if i = 0 then lst else drop' 1 lst - -let take i lst = - let rec take' i' lst' = - match lst' with - | l :: ls -> if i' = i then [ l ] else l :: take' (i' + 1) ls - | [] -> [] - in - if i = 0 then [] else take' 1 lst - -let rec next_store i = function - | [] -> None - | RBstore (_, _, _, _, _) :: _ -> Some i - | _ :: rst -> next_store (i + 1) rst - -let rec next_load i = function - | [] -> None - | RBload (_, _, _, _, _) :: _ -> Some i - | _ :: rst -> next_load (i + 1) rst - -let accumulate_RAW_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBload (_, _, _, _, _) -> ( - 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) ) - | _ -> dfg - -let accumulate_WAR_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBstore (_, _, _, _, _) -> ( - match next_load 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) ) - | _ -> dfg - -let accumulate_WAW_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | 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)) - | _ -> dfg - -(** Predicate dependencies. *) - -let rec in_predicate p p' = - match p' with - | Plit p'' -> P.to_int p = P.to_int (snd p'') - | Pand (p1, p2) -> in_predicate p p1 || in_predicate p p2 - | Por (p1, p2) -> in_predicate p p1 || in_predicate p p2 - | Ptrue -> false - | Pfalse -> false - -let get_predicate = function - | RBop (p, _, _, _) -> p - | RBload (p, _, _, _, _) -> p - | RBstore (p, _, _, _, _) -> p - | RBsetpred (p, _, _, _) -> p - | _ -> None - -let rec next_setpred p i = function - | [] -> None - | RBsetpred (_, _, _, p') :: rst -> - if in_predicate p' p then - Some i - else - next_setpred p (i + 1) rst - | _ :: rst -> next_setpred p (i + 1) rst - -let rec next_preduse p i instr= - let next p' rst = - if in_predicate p p' then - Some i - else - next_preduse p (i + 1) rst - in - match instr with - | [] -> None - | 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 = - let i, curr = curri in - match get_predicate curr with - | Some p -> ( - match next_setpred p 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) ) - | _ -> dfg - -let accumulate_WAR_pred_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBsetpred (_, _, _, p) -> ( - match next_preduse p 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) ) - | _ -> dfg - -let accumulate_WAW_pred_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBsetpred (_, _, _, p) -> ( - match next_setpred (Plit (true, p)) 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) ) - | _ -> dfg - -(** This function calculates the WAW dependencies, which happen when two writes are ordered one - after another and therefore have to be kept in that order. This accumulation might be redundant - if register renaming is done before hand, because then these dependencies can be avoided. *) -let accumulate_WAW_deps instrs dfg curri = - let i, curr = curri in - let dst_dep dst = - match find_next_dst_write i dst (i + 1) (drop (i + 1) instrs) with - | Some (a, b) -> DFG.add_edge dfg (gen_vertex instrs a) (gen_vertex instrs b) - | _ -> dfg - in - match curr with - | RBop (_, _, _, dst) -> dst_dep dst - | RBload (_, _, _, _, dst) -> dst_dep dst - | RBstore (_, _, _, _, _) -> ( - match next_store (i + 1) (drop (i + 1) instrs) with - | None -> dfg - | Some i' -> DFG.add_edge dfg (gen_vertex instrs i) (gen_vertex instrs i') ) - | _ -> dfg - -let accumulate_WAR_deps instrs dfg curri = - let i, curr = curri in - let dst_dep dst = - let dep_list = find_all_next_dst_read i dst 0 (take i instrs |> List.rev) - |> List.map (function (d, d') -> (i - d' - 1, d)) - in - List.fold_left (fun g -> - function (d, d') -> DFG.add_edge g (gen_vertex instrs d) (gen_vertex instrs d')) dfg dep_list - in - match curr with - | RBop (_, _, _, dst) -> dst_dep dst - | RBload (_, _, _, _, dst) -> dst_dep dst - | _ -> dfg - -let assigned_vars vars = function - | RBnop -> vars - | RBop (_, _, _, dst) -> dst :: vars - | RBload (_, _, _, _, dst) -> dst :: vars - | RBstore (_, _, _, _, _) -> vars - | RBsetpred (_, _, _, _) -> vars - -let get_pred = function - | RBnop -> None - | RBop (op, _, _, _) -> op - | RBload (op, _, _, _, _) -> op - | RBstore (op, _, _, _, _) -> op - | RBsetpred (_, _, _, _) -> None - -let independant_pred p p' = - match sat_pred_simple (Pand (p, p')) with - | None -> true - | _ -> false - -let check_dependent op1 op2 = - match op1, op2 with - | Some p, Some p' -> not (independant_pred p p') - | _, _ -> true - -let remove_unnecessary_deps graph = - let is_dependent v1 v2 g = - let (_, instr1) = v1 in - let (_, instr2) = v2 in - if check_dependent (get_pred instr1) (get_pred instr2) - then g - else DFG.remove_edge g v1 v2 - in - DFG.fold_edges is_dependent graph graph - -(** All the nodes in the DFG have to come after the source of the basic block, and should terminate - before the sink of the basic block. After that, there should be constraints for data - dependencies between nodes. *) -let gather_bb_constraints debug bb = - let ibody = List.mapi (fun i a -> (i, a)) bb.bb_body in - let dfg = List.fold_left (fun dfg v -> DFG.add_vertex dfg v) DFG.empty ibody in - let _, _, dfg' = - List.fold_left (accumulate_RAW_deps ibody) - (0, PTree.empty, dfg) - bb.bb_body - in - let dfg'' = List.fold_left (fun dfg f -> List.fold_left (f bb.bb_body) dfg ibody) dfg' - [ accumulate_WAW_deps; - accumulate_WAR_deps; - accumulate_RAW_mem_deps; - accumulate_WAR_mem_deps; - accumulate_WAW_mem_deps; - accumulate_RAW_pred_deps; - accumulate_WAR_pred_deps; - accumulate_WAW_pred_deps - ] - in - let dfg''' = remove_unnecessary_deps dfg'' in - (List.length bb.bb_body, dfg''', successors_instr bb.bb_exit) - -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 - then G.add_edge_e g (encode_bb n 0, 0, encode_var n (fst v1) 0) - else g) |> - (fun g' -> - if DFG.out_degree dfg v1 = 0 - 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, 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 = - List.fold_left (fun g n' -> - G.add_edge_e g (encode_bb n 1, -1, encode_bb n' 0) - ) constr succs - -module BFDFG = Graph.Path.BellmanFord(DFG)(struct - include DFG - type t = int - let weight = DFG.E.label - let compare = compare - let add = (+) - let zero = 0 - end) - -module TopoDFG = Graph.Topological.Make(DFG) - -let negate_graph constr = - DFG.fold_edges_e (function (v1, e, v2) -> fun g -> - DFG.add_edge_e g (v1, -e, v2) - ) constr DFG.empty - -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 - |> 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 maxf))) - 1), - encode_var n (fst v') 0) - ) constr (DFG.fold_vertex (fun v l -> - 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 = - | Mem - | SDiv - | UDiv - -type resources = { - res_mem: DFG.V.t list; - res_udiv: DFG.V.t list; - res_sdiv: DFG.V.t list; -} - -let find_resource = function - | RBload _ -> Some Mem - | RBstore _ -> Some Mem - | RBop (_, op, _, _) -> - ( match op with - | Odiv -> Some SDiv - | Odivu -> Some UDiv - | Omod -> Some SDiv - | Omodu -> Some UDiv - | _ -> None ) - | _ -> None - -let add_resource_constr n dfg constr = - let res = TopoDFG.fold (function (i, instr) -> - function {res_mem = ml; res_sdiv = sdl; res_udiv = udl} as r -> - match find_resource instr with - | Some SDiv -> {r with res_sdiv = (i, instr) :: sdl} - | Some UDiv -> {r with res_udiv = (i, instr) :: udl} - | Some Mem -> {r with res_mem = (i, instr) :: ml} - | None -> r - ) dfg {res_mem = []; res_sdiv = []; res_udiv = []} - in - let get_constraints l g = List.fold_left (fun gv v' -> - match gv with - | (g, None) -> (g, Some v') - | (g, Some v) -> - (G.add_edge_e g (encode_var n (fst v) 0, -1, encode_var n (fst v') 0), Some v') - ) (g, None) l |> fst - in - get_constraints (List.rev res.res_mem) constr - |> get_constraints (List.rev res.res_udiv) - |> get_constraints (List.rev res.res_sdiv) - -let gather_cfg_constraints c constr curr = - let (n, dfg) = curr in - match PTree.get (P.of_int n) c with - | 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 - |> List.filter (fun n' -> n' < n)) - |> add_cycle_constr 8 n dfg - |> add_resource_constr n dfg - -let rec intersperse s = function - | [] -> [] - | [ a ] -> [ a ] - | x :: xs -> x :: s :: intersperse s xs - -let print_objective constr = - let vars = G.fold_vertex (fun v1 l -> - match v1 with - | { sv_type = VarType _; sv_num = 0 } -> print_sv v1 :: l - | _ -> l - ) constr [] - in - "min: " ^ String.concat "" (intersperse " + " vars) ^ ";\n" - -let print_lp constr = - print_objective constr ^ - (G.fold_edges_e (function (e1, w, e2) -> fun s -> - s ^ sprintf "%s - %s <= %d;\n" (print_sv e1) (print_sv e2) w - ) constr "" |> - G.fold_vertex (fun v1 s -> - s ^ sprintf "int %s;\n" (print_sv v1) - ) constr) - -let update_schedule v = function Some l -> Some (v :: l) | None -> Some [ v ] - -let parse_soln (tree, bbtree) s = - let r = Str.regexp "var\\([0-9]+\\)n\\([0-9]+\\)_0[ ]+\\([0-9]+\\)" in - 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 (fn, oc) = Filename.open_temp_file "vericert_" "_lp_solve" in - fprintf oc "%s\n" (print_lp constr); - close_out oc; - - let res = Str.split (Str.regexp_string "\n") (read_process ("lp_solve " ^ fn)) - |> drop 3 - |> List.fold_left parse_soln (IMap.empty, 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 - List.fold_left (fun g v -> - List.fold_left (fun g' v' -> - let edges = DFG.find_all_edges dfg v v' in - List.fold_left DFG.add_edge_e g' edges - ) g l - ) dfg' l - -let rec all_successors dfg v = - List.concat (List.fold_left (fun l v -> - all_successors dfg v :: l - ) [] (DFG.succ dfg v)) - -let order_instr dfg = - DFG.fold_vertex (fun v li -> - if DFG.in_degree dfg v = 0 - then (List.map snd (v :: all_successors dfg v)) :: li - else li - ) dfg [] - -let combine_bb_schedule schedule s = - let i, st = s in - IMap.update st (update_schedule i) schedule - -(**let add_el dfg i l = - List.*) - -let check_in el = - List.exists (List.exists ((=) el)) - -let all_dfs dfg = - let roots = DFG.fold_vertex (fun v li -> - if DFG.in_degree dfg v = 0 then v :: li else li - ) dfg [] in - let dfg' = DFG.fold_edges (fun v1 v2 g -> DFG.add_edge g v2 v1) dfg dfg in - List.fold_left (fun a el -> - if check_in el a then a else - (DFGDfs.fold_component (fun v l -> v :: l) [] dfg' el) :: a) [] roots - -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 = - 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 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 - |> 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 -> - all_dfs x - |> List.map (subgraph x) - |> List.map (fun y -> - TopoDFG.fold (fun i l -> snd i :: l) y [] - |> List.rev))) body2 - (*|> (fun x -> TopoDFG.fold (fun i l -> snd i :: l) x []) - |> List.rev) body2*) - in - { bb_body = final_body2; - bb_exit = ctrl_flow - } - in - PTree.map f c - -let schedule entry (c : RTLBlock.bb RTLBlockInstr.code) = - let debug = true in - let transf_graph (_, dfg, _) = dfg in - let c' = PTree.map1 (fun x -> gather_bb_constraints false x |> transf_graph) c in - (*let _ = if debug then PTree.map (fun r o -> printf "##### %d #####\n%a\n\n" (P.to_int r) print_dfg o) c' else PTree.empty in*) - let cgraph = PTree.elements c' - |> List.map (function (x, y) -> (P.to_int x, y)) - |> List.fold_left (gather_cfg_constraints c) G.empty - in - let graph = open_out "constr_graph.dot" in - fprintf graph "%a\n" GDot.output_graph cgraph; - close_out graph; - let schedule' = solve_constraints cgraph in - (**IMap.iter (fun a b -> printf "##### %d #####\n%a\n\n" a (print_list print_tuple) b) schedule';*) - (**printf "Schedule: %a\n" (fun a x -> IMap.iter (fun d -> fprintf a "%d: %a\n" d (print_list print_tuple)) x) schedule';*) - transf_rtlpar c c' schedule' - -let rec find_reachable_states c e = - match PTree.get e c with - | Some { bb_exit = ex; _ } -> - e :: List.fold_left (fun x a -> List.concat [x; find_reachable_states c a]) [] - (successors_instr ex |> List.filter (fun x -> P.lt x e)) - | None -> assert false - -let add_to_tree c nt i = - match PTree.get i c with - | Some p -> PTree.set i p nt - | None -> assert false - -let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = - let scheduled = schedule f.fn_entrypoint f.fn_code in - let reachable = find_reachable_states scheduled f.fn_entrypoint - |> List.to_seq |> SS.of_seq |> SS.to_seq |> List.of_seq in - { fn_sig = f.fn_sig; - fn_params = f.fn_params; - fn_stacksize = f.fn_stacksize; - fn_code = scheduled (*List.fold_left (add_to_tree scheduled) PTree.empty reachable*); - fn_entrypoint = f.fn_entrypoint - } -#+end_src - -* RTLPargen -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargen.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpargen-main -#+begin_src coq -Require Import compcert.backend.Registers. -Require Import compcert.common.AST. -Require Import compcert.common.Globalenvs. -Require Import compcert.common.Memory. -Require Import compcert.common.Values. -Require Import compcert.lib.Floats. -Require Import compcert.lib.Integers. -Require Import compcert.lib.Maps. -Require compcert.verilog.Op. - -Require Import vericert.common.Vericertlib. -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLPar. -Require Import vericert.hls.RTLBlockInstr. -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. - -(** ** 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) - end. - -Fixpoint replicate {A} (n: nat) (l: A) := - match n with - | O => nil - | S n => l :: replicate n l - end. - -Definition merge''' x y := - match x, y with - | Some p1, Some p2 => Some (Pand p1 p2) - | Some p, None | None, Some p => Some p - | None, None => None - end. - -Definition merge'' x := - match x with - | ((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 := - match pa, pf with - | (p, a), (p', f) => (merge''' p p', f a) - end. - -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. - -(*map (fun x => (fst x, Econs (snd x) Enil)) pel*) -Definition merge' (pel: pred_expr) (tpel: predicated expression_list) := - predicated_map (uncurry Econs) (predicated_prod pel tpel). - -Fixpoint merge (pel: list pred_expr): predicated expression_list := - match pel with - | nil => NE.singleton (T, Enil) - | a :: b => merge' a (merge b) - end. - -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 := - 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 := - 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 := - NE.map (fun x => (fst x, (snd x) pa pb pc)) pf. - -Definition predicated_from_opt {A: Type} (p: option pred_op) (a: A) := - match p with - | Some p' => NE.singleton (p', a) - | None => NE.singleton (T, a) - end. - -#[local] Open Scope non_empty_scope. -#[local] Open Scope pred_op. - -Fixpoint NEfold_left {A B} (f: A -> B -> A) (l: NE.non_empty B) (a: A) : A := - match l with - | NE.singleton a' => f a a' - | a' ::| b => NEfold_left f b (f a a') - end. - -Fixpoint NEapp {A} (l m: NE.non_empty A) := - match l with - | NE.singleton a => a ::| m - | a ::| b => a ::| NEapp b m - end. - -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. - -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 - -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 - | RBop p op rl r => - f # (Reg r) <- - (app_predicated p - (f # (Reg r)) - (map_predicated (pred_ret (Eop op)) (merge (list_translation rl f)))) - | RBload p chunk addr rl r => - f # (Reg r) <- - (app_predicated p - (f # (Reg r)) - (map_predicated - (map_predicated (pred_ret (Eload chunk addr)) (merge (list_translation rl f))) - (f # Mem))) - | RBstore p chunk addr rl r => - f # Mem <- - (app_predicated p - (f # Mem) - (map_predicated - (map_predicated - (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)))) - 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. - -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. - -(** 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. *) - -Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := - match bb with - | nil => - match bbt with - | nil => true - | _ => false - end - | _ => true - end. - -Definition schedule_oracle (bb: RTLBlock.bblock) (bbt: RTLPar.bblock) : bool := - check (abstract_sequence empty (bb_body bb)) - (abstract_sequence empty (concat (concat (bb_body bbt)))) && - check_control_flow_instr (bb_exit bb) (bb_exit bbt) && - empty_trees (bb_body bb) (bb_body bbt). - -Definition check_scheduled_trees := beq2 schedule_oracle. - -Ltac solve_scheduled_trees_correct := - intros; unfold check_scheduled_trees in *; - match goal with - | [ H: context[beq2 _ _ _], x: positive |- _ ] => - rewrite beq2_correct in H; specialize (H x) - end; repeat destruct_match; crush. - -Lemma check_scheduled_trees_correct: - forall f1 f2 x y1, - check_scheduled_trees f1 f2 = true -> - PTree.get x f1 = Some y1 -> - exists y2, PTree.get x f2 = Some y2 /\ schedule_oracle y1 y2 = true. -Proof. solve_scheduled_trees_correct; eexists; crush. Qed. - -Lemma check_scheduled_trees_correct2: - forall f1 f2 x, - check_scheduled_trees f1 f2 = true -> - PTree.get x f1 = None -> - PTree.get x f2 = None. -Proof. solve_scheduled_trees_correct. Qed. - -(** ** Top-level Functions *) - -Parameter schedule : RTLBlock.function -> 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) - f.(fn_params) - f.(fn_stacksize) - tfcode - f.(fn_entrypoint)) - else - 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. -#+end_src -* RTLPargenproof -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargenproof.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpargenproof-imports -#+begin_src coq -Require Import compcert.backend.Registers. -Require Import compcert.common.AST. -Require Import compcert.common.Errors. -Require Import compcert.common.Linking. -Require Import compcert.common.Globalenvs. -Require Import compcert.common.Memory. -Require Import compcert.common.Values. -Require Import compcert.lib.Maps. - -Require Import vericert.common.Vericertlib. -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLPar. -Require Import vericert.hls.RTLBlockInstr. -Require Import vericert.hls.RTLPargen. -Require Import vericert.hls.Predicate. -Require Import vericert.hls.Abstr. - -#[local] Open Scope positive. -#[local] Open Scope forest. -#[local] Open Scope pred_op. -#+end_src - -** RTLBlock to abstract translation - -Correctness of translation from RTLBlock to the abstract interpretation language. - -#+name: rtlpargenproof-main -#+begin_src coq -(*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. - -Inductive state_lessdef : instr_state -> instr_state -> Prop := - state_lessdef_intro : - forall rs1 rs2 m1, - (forall x, rs1 !! x = rs2 !! x) -> - state_lessdef (mk_instr_state rs1 m1) (mk_instr_state rs2 m1). - -Ltac inv_simp := - repeat match goal with - | H: exists _, _ |- _ => inv H - end; simplify. - -*) - -Definition check_dest i r' := - match i with - | RBop p op rl r => (r =? r')%positive - | RBload p chunk addr rl r => (r =? r')%positive - | _ => false - end. - -Lemma check_dest_dec i r : {check_dest i r = true} + {check_dest i r = false}. -Proof. destruct (check_dest i r); tauto. Qed. - -Fixpoint check_dest_l l r := - match l with - | nil => false - | a :: b => check_dest a r || check_dest_l b r - end. - -Lemma check_dest_l_forall : - forall l r, - check_dest_l l r = false -> - Forall (fun x => check_dest x r = false) l. -Proof. induction l; crush. Qed. - -Lemma check_dest_l_dec i r : {check_dest_l i r = true} + {check_dest_l i r = false}. -Proof. destruct (check_dest_l i r); tauto. Qed. - -Lemma check_dest_update : - forall f i r, - check_dest i r = false -> - (update f i) # (Reg r) = f # (Reg r). -Proof. - destruct i; crush; try apply Pos.eqb_neq in H; apply genmap1; crush. -Qed. - -Lemma check_dest_l_forall2 : - forall l r, - Forall (fun x => check_dest x r = false) l -> - check_dest_l l r = false. -Proof. - induction l; crush. - inv H. apply orb_false_intro; crush. -Qed. - -Lemma check_dest_l_ex2 : - forall l r, - (exists a, In a l /\ check_dest a r = true) -> - check_dest_l l r = true. -Proof. - induction l; crush. - specialize (IHl r). inv H. - apply orb_true_intro; crush. - apply orb_true_intro; crush. - right. apply IHl. exists x. auto. -Qed. - -Lemma check_list_l_false : - forall l x r, - check_dest_l (l ++ x :: nil) r = false -> - check_dest_l l r = false /\ check_dest x r = false. -Proof. - simplify. - apply check_dest_l_forall in H. apply Forall_app in H. - simplify. apply check_dest_l_forall2; auto. - apply check_dest_l_forall in H. apply Forall_app in H. - simplify. inv H1. auto. -Qed. - -Lemma check_dest_l_ex : - forall l r, - check_dest_l l r = true -> - exists a, In a l /\ check_dest a r = true. -Proof. - induction l; crush. - destruct (check_dest a r) eqn:?; try solve [econstructor; crush]. - simplify. - exploit IHl. apply H. simplify. econstructor. simplify. right. eassumption. - auto. -Qed. - -Lemma check_list_l_true : - forall l x r, - check_dest_l (l ++ x :: nil) r = true -> - check_dest_l l r = true \/ check_dest x r = true. -Proof. - simplify. - apply check_dest_l_ex in H; simplify. - apply in_app_or in H. inv H. left. - apply check_dest_l_ex2. exists x0. auto. - inv H0; auto. -Qed. - -Lemma check_dest_l_dec2 l r : - {Forall (fun x => check_dest x r = false) l} - + {exists a, In a l /\ check_dest a r = true}. -Proof. - destruct (check_dest_l_dec l r); [right | left]; - auto using check_dest_l_ex, check_dest_l_forall. -Qed. - -Lemma abstr_comp : - forall l i f x x0, - abstract_sequence f (l ++ i :: nil) = x -> - abstract_sequence f l = x0 -> - x = update x0 i. -Proof. induction l; intros; crush; eapply IHl; eauto. Qed. - -(* - -Lemma gen_list_base: - forall FF ge sp l rs exps st1, - (forall x, @sem_value FF ge sp st1 (exps # (Reg x)) (rs !! x)) -> - sem_val_list ge sp st1 (list_translation l exps) rs ## l. -Proof. - induction l. - intros. simpl. constructor. - intros. simpl. eapply Scons; eauto. -Qed. - -Lemma check_dest_update2 : - forall f r rl op p, - (update f (RBop p op rl r)) # (Reg r) = Eop op (list_translation rl f) (f # Mem). -Proof. crush; rewrite map2; auto. Qed. - -Lemma check_dest_update3 : - forall f r rl p addr chunk, - (update f (RBload p chunk addr rl r)) # (Reg r) = Eload chunk addr (list_translation rl f) (f # Mem). -Proof. crush; rewrite map2; auto. Qed. - -Lemma abstract_seq_correct_aux: - forall FF ge sp i st1 st2 st3 f, - @step_instr FF ge sp st3 i st2 -> - sem ge sp st1 f st3 -> - sem ge sp st1 (update f i) st2. -Proof. - intros; inv H; simplify. - { simplify; eauto. } (*apply match_states_refl. }*) - { inv H0. inv H6. destruct st1. econstructor. simplify. - constructor. intros. - destruct (resource_eq (Reg res) (Reg x)). inv e. - rewrite map2. econstructor. eassumption. apply gen_list_base; eauto. - rewrite Regmap.gss. eauto. - assert (res <> x). { unfold not in *. intros. apply n. rewrite H0. auto. } - rewrite Regmap.gso by auto. - rewrite genmap1 by auto. auto. - - rewrite genmap1; crush. } - { inv H0. inv H7. constructor. constructor. intros. - destruct (Pos.eq_dec dst x); subst. - rewrite map2. econstructor; eauto. - apply gen_list_base. auto. rewrite Regmap.gss. auto. - rewrite genmap1. rewrite Regmap.gso by auto. auto. - unfold not in *; intros. inv H0. auto. - rewrite genmap1; crush. - } - { inv H0. inv H7. constructor. constructor; intros. - rewrite genmap1; crush. - rewrite map2. econstructor; eauto. - apply gen_list_base; auto. - } -Qed. - -Lemma regmap_list_equiv : - forall A (rs1: Regmap.t A) rs2, - (forall x, rs1 !! x = rs2 !! x) -> - forall rl, rs1##rl = rs2##rl. -Proof. induction rl; crush. Qed. - -Lemma sem_update_Op : - forall A ge sp st f st' r l o0 o m rs v, - @sem A ge sp st f st' -> - Op.eval_operation ge sp o0 rs ## l m = Some v -> - match_states st' (mk_instr_state rs m) -> - exists tst, - sem ge sp st (update f (RBop o o0 l r)) tst /\ match_states (mk_instr_state (Regmap.set r v rs) m) tst. -Proof. - intros. inv H1. simplify. - destruct st. - econstructor. simplify. - { constructor. - { constructor. intros. destruct (Pos.eq_dec x r); subst. - { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. - { inv H9. eapply gen_list_base; eauto. } - { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } - { inv H. rewrite genmap1; crush. eauto. } } - { constructor; eauto. intros. - destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed. - -Lemma sem_update_load : - forall A ge sp st f st' r o m a l m0 rs v a0, - @sem A ge sp st f st' -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.loadv m m0 a0 = Some v -> - match_states st' (mk_instr_state rs m0) -> - exists tst : instr_state, - sem ge sp st (update f (RBload o m a l r)) tst - /\ match_states (mk_instr_state (Regmap.set r v rs) m0) tst. -Proof. - intros. inv H2. pose proof H. inv H. inv H9. - destruct st. - econstructor; simplify. - { constructor. - { constructor. intros. - destruct (Pos.eq_dec x r); subst. - { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. - rewrite <- H6. eauto. - instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. auto. } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. } } - { rewrite genmap1; crush. eauto. } } - { constructor; auto; intros. destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed. - -Lemma sem_update_store : - forall A ge sp a0 m a l r o f st m' rs m0 st', - @sem A ge sp st f st' -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.storev m m0 a0 rs !! r = Some m' -> - match_states st' (mk_instr_state rs m0) -> - exists tst, sem ge sp st (update f (RBstore o m a l r)) tst - /\ match_states (mk_instr_state rs m') tst. -Proof. - intros. inv H2. pose proof H. inv H. inv H9. - destruct st. - econstructor; simplify. - { econstructor. - { econstructor; intros. rewrite genmap1; crush. } - { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. rewrite <- H6. - eauto. specialize (H6 r). rewrite H6. eauto. } } - { econstructor; eauto. } -Qed. - -Lemma sem_update : - forall A ge sp st x st' st'' st''' f, - sem ge sp st f st' -> - match_states st' st''' -> - @step_instr A ge sp st''' x st'' -> - exists tst, sem ge sp st (update f x) tst /\ match_states st'' tst. -Proof. - intros. destruct x; inv H1. - { econstructor. split. - apply sem_update_RBnop. eassumption. - apply match_states_commut. auto. } - { eapply sem_update_Op; eauto. } - { eapply sem_update_load; eauto. } - { eapply sem_update_store; eauto. } -Qed. - -Lemma sem_update2_Op : - forall A ge sp st f r l o0 o m rs v, - @sem A ge sp st f (mk_instr_state rs m) -> - Op.eval_operation ge sp o0 rs ## l m = Some v -> - sem ge sp st (update f (RBop o o0 l r)) (mk_instr_state (Regmap.set r v rs) m). -Proof. - intros. destruct st. constructor. - inv H. inv H6. - { constructor; intros. simplify. - destruct (Pos.eq_dec r x); subst. - { rewrite map2. econstructor. eauto. - apply gen_list_base. eauto. - rewrite Regmap.gss. auto. } - { rewrite genmap1; crush. rewrite Regmap.gso; auto. } } - { simplify. rewrite genmap1; crush. inv H. eauto. } -Qed. - -Lemma sem_update2_load : - forall A ge sp st f r o m a l m0 rs v a0, - @sem A ge sp st f (mk_instr_state rs m0) -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.loadv m m0 a0 = Some v -> - sem ge sp st (update f (RBload o m a l r)) (mk_instr_state (Regmap.set r v rs) m0). -Proof. - intros. simplify. inv H. inv H7. constructor. - { constructor; intros. destruct (Pos.eq_dec r x); subst. - { rewrite map2. rewrite Regmap.gss. econstructor; eauto. - apply gen_list_base; eauto. } - { rewrite genmap1; crush. rewrite Regmap.gso; eauto. } - } - { simplify. rewrite genmap1; crush. } -Qed. - -Lemma sem_update2_store : - forall A ge sp a0 m a l r o f st m' rs m0, - @sem A ge sp st f (mk_instr_state rs m0) -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.storev m m0 a0 rs !! r = Some m' -> - sem ge sp st (update f (RBstore o m a l r)) (mk_instr_state rs m'). -Proof. - intros. simplify. inv H. inv H7. constructor; simplify. - { econstructor; intros. rewrite genmap1; crush. } - { rewrite map2. econstructor; eauto. apply gen_list_base; eauto. } -Qed. - -Lemma sem_update2 : - forall A ge sp st x st' st'' f, - sem ge sp st f st' -> - @step_instr A ge sp st' x st'' -> - sem ge sp st (update f x) st''. -Proof. - intros. - destruct x; inv H0; - eauto using sem_update_RBnop, sem_update2_Op, sem_update2_load, sem_update2_store. -Qed. - -Lemma abstr_sem_val_mem : - forall A B ge tge st tst sp a, - ge_preserved ge tge -> - forall v m, - (@sem_mem A ge sp st a m /\ match_states st tst -> @sem_mem B tge sp tst a m) /\ - (@sem_value A ge sp st a v /\ match_states st tst -> @sem_value B tge sp tst a v). -Proof. - intros * H. - apply expression_ind2 with - - (P := fun (e1: expression) => - forall v m, - (@sem_mem A ge sp st e1 m /\ match_states st tst -> @sem_mem B tge sp tst e1 m) /\ - (@sem_value A ge sp st e1 v /\ match_states st tst -> @sem_value B tge sp tst e1 v)) - - (P0 := fun (e1: expression_list) => - forall lv, @sem_val_list A ge sp st e1 lv /\ match_states st tst -> @sem_val_list B tge sp tst e1 lv); - simplify; intros; simplify. - { inv H1. inv H2. constructor. } - { inv H2. inv H1. rewrite H0. constructor. } - { inv H3. } - { inv H3. inv H4. econstructor. apply H1; auto. simplify. eauto. constructor. auto. auto. - apply H0; simplify; eauto. constructor; eauto. - unfold ge_preserved in *. simplify. rewrite <- H2. auto. - } - { inv H3. } - { inv H3. inv H4. econstructor. apply H1; eauto; simplify; eauto. constructor; eauto. - apply H0; simplify; eauto. constructor; eauto. - inv H. rewrite <- H4. eauto. - auto. - } - { inv H4. inv H5. econstructor. apply H0; eauto. simplify; eauto. constructor; eauto. - apply H2; eauto. simplify; eauto. constructor; eauto. - apply H1; eauto. simplify; eauto. constructor; eauto. - inv H. rewrite <- H5. eauto. auto. - } - { inv H4. } - { inv H1. constructor. } - { inv H3. constructor; auto. apply H0; eauto. apply Mem.empty. } -Qed. - -Lemma abstr_sem_value : - forall a A B ge tge sp st tst v, - @sem_value A ge sp st a v -> - ge_preserved ge tge -> - match_states st tst -> - @sem_value B tge sp tst a v. -Proof. intros; eapply abstr_sem_val_mem; eauto; apply Mem.empty. Qed. - -Lemma abstr_sem_mem : - forall a A B ge tge sp st tst v, - @sem_mem A ge sp st a v -> - ge_preserved ge tge -> - match_states st tst -> - @sem_mem B tge sp tst a v. -Proof. intros; eapply abstr_sem_val_mem; eauto. Qed. - -Lemma abstr_sem_regset : - forall a a' A B ge tge sp st tst rs, - @sem_regset A ge sp st a rs -> - ge_preserved ge tge -> - (forall x, a # x = a' # x) -> - match_states st tst -> - exists rs', @sem_regset B tge sp tst a' rs' /\ (forall x, rs !! x = rs' !! x). -Proof. - inversion 1; intros. - inv H7. - econstructor. simplify. econstructor. intros. - eapply abstr_sem_value; eauto. rewrite <- H6. - eapply H0. constructor; eauto. - auto. -Qed. - -Lemma abstr_sem : - forall a a' A B ge tge sp st tst st', - @sem A ge sp st a st' -> - ge_preserved ge tge -> - (forall x, a # x = a' # x) -> - match_states st tst -> - exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. -Proof. - inversion 1; subst; intros. - inversion H4; subst. - exploit abstr_sem_regset; eauto; inv_simp. - do 3 econstructor; eauto. - rewrite <- H3. - eapply abstr_sem_mem; eauto. -Qed. - -Lemma abstract_execution_correct': - forall A B ge tge sp st' a a' st tst, - @sem A ge sp st a st' -> - ge_preserved ge tge -> - check a a' = true -> - match_states st tst -> - exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. -Proof. - intros; - pose proof (check_correct a a' H1); - eapply abstr_sem; eauto. -Qed. - -Lemma states_match : - forall st1 st2 st3 st4, - match_states st1 st2 -> - match_states st2 st3 -> - match_states st3 st4 -> - match_states st1 st4. -Proof. - intros * H1 H2 H3; destruct st1; destruct st2; destruct st3; destruct st4. - inv H1. inv H2. inv H3; constructor. - unfold regs_lessdef in *. intros. - repeat match goal with - | H: forall _, _, r : positive |- _ => specialize (H r) - end. - congruence. - auto. -Qed. - -Lemma step_instr_block_same : - forall ge sp st st', - step_instr_block ge sp st nil st' -> - st = st'. -Proof. inversion 1; auto. Qed. - -Lemma step_instr_seq_same : - forall ge sp st st', - step_instr_seq ge sp st nil st' -> - st = st'. -Proof. inversion 1; auto. Qed. - -Lemma sem_update' : - forall A ge sp st a x st', - sem ge sp st (update (abstract_sequence empty a) x) st' -> - exists st'', - @step_instr A ge sp st'' x st' /\ - sem ge sp st (abstract_sequence empty a) st''. -Proof. - Admitted. - -Lemma rtlpar_trans_correct : - forall bb ge sp sem_st' sem_st st, - sem ge sp sem_st (abstract_sequence empty (concat (concat bb))) sem_st' -> - match_states sem_st st -> - exists st', RTLPar.step_instr_block ge sp st bb st' - /\ match_states sem_st' st'. -Proof. - induction bb using rev_ind. - { repeat econstructor. eapply abstract_interp_empty3 in H. - inv H. inv H0. constructor; congruence. } - { simplify. inv H0. repeat rewrite concat_app in H. simplify. - rewrite app_nil_r in H. - exploit sem_separate; eauto; inv_simp. - repeat econstructor. admit. admit. - } -Admitted. - -(*Lemma abstract_execution_correct_ld: - forall bb bb' cfi ge tge sp st st' tst, - RTLBlock.step_instr_list ge sp st bb st' -> - ge_preserved ge tge -> - schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi) = true -> - match_states_ld st tst -> - exists tst', RTLPar.step_instr_block tge sp tst bb' tst' - /\ match_states st' tst'. -Proof. - intros.*) -*) - -Lemma match_states_list : - forall A (rs: Regmap.t A) rs', - (forall r, rs !! r = rs' !! r) -> - forall l, rs ## l = rs' ## l. -Proof. induction l; crush. Qed. - -Lemma PTree_matches : - forall A (v: A) res rs rs', - (forall r, rs !! r = rs' !! r) -> - forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x. -Proof. - intros; destruct (Pos.eq_dec x res); subst; - [ repeat rewrite Regmap.gss by auto - | repeat rewrite Regmap.gso by auto ]; auto. -Qed. - -Lemma abstract_interp_empty3 : - forall A ctx st', - @sem A ctx empty st' -> match_states (ctx_is ctx) st'. -Proof. - inversion 1; subst; simplify. destruct ctx. - destruct ctx_is. - constructor; intros. - - inv H0. specialize (H3 x). inv H3. inv H8. reflexivity. - - inv H1. specialize (H3 x). inv H3. inv H8. reflexivity. - - inv H2. inv H8. reflexivity. -Qed. - -Lemma step_instr_matches : - forall A a ge sp st st', - @step_instr A ge sp st a st' -> - forall tst, - match_states st tst -> - exists tst', step_instr ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction 1; simplify; - match goal with H: match_states _ _ |- _ => inv H end; - try solve [repeat econstructor; try erewrite match_states_list; - try apply PTree_matches; eauto; - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end]. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - apply PTree_matches; assumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - econstructor; auto. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - apply PTree_matches; assumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - econstructor; auto. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - erewrite <- eval_predf_pr_equiv; eassumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - erewrite <- eval_predf_pr_equiv; eassumption. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - - admit. Admitted. - -Lemma step_instr_list_matches : - forall a ge sp st st', - step_instr_list ge sp st a st' -> - forall tst, match_states st tst -> - exists tst', step_instr_list ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction a; intros; inv H; - try (exploit step_instr_matches; eauto; []; simplify; - exploit IHa; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma step_instr_seq_matches : - forall a ge sp st st', - step_instr_seq ge sp st a st' -> - forall tst, match_states st tst -> - exists tst', step_instr_seq ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction a; intros; inv H; - try (exploit step_instr_list_matches; eauto; []; simplify; - exploit IHa; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma step_instr_block_matches : - forall bb ge sp st st', - step_instr_block ge sp st bb st' -> - forall tst, match_states st tst -> - exists tst', step_instr_block ge sp tst bb tst' - /\ match_states st' tst'. -Proof. - induction bb; intros; inv H; - try (exploit step_instr_seq_matches; eauto; []; simplify; - exploit IHbb; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma rtlblock_trans_correct' : - forall bb ge sp st x st'', - RTLBlock.step_instr_list ge sp st (bb ++ x :: nil) st'' -> - exists st', RTLBlock.step_instr_list ge sp st bb st' - /\ step_instr ge sp st' x st''. -Proof. - induction bb. - crush. exists st. - split. constructor. inv H. inv H6. auto. - crush. inv H. exploit IHbb. eassumption. simplify. - econstructor. split. - econstructor; eauto. eauto. -Qed. - -Lemma abstract_interp_empty A st : @sem A st empty (ctx_is st). -Proof. destruct st, ctx_is. simpl. repeat econstructor. Qed. - -Lemma abstract_seq : - forall l f i, - abstract_sequence f (l ++ i :: nil) = update (abstract_sequence f l) i. -Proof. induction l; crush. Qed. - -Lemma abstract_sequence_update : - forall l r f, - check_dest_l l r = false -> - (abstract_sequence f l) # (Reg r) = f # (Reg r). -Proof. - induction l using rev_ind; crush. - rewrite abstract_seq. rewrite check_dest_update. apply IHl. - apply check_list_l_false in H. tauto. - apply check_list_l_false in H. tauto. -Qed. - -(*Lemma sem_separate : - forall A ctx b a st', - sem ctx (abstract_sequence empty (a ++ b)) st' -> - exists st'', - @sem A ctx (abstract_sequence empty a) st'' - /\ @sem A (mk_ctx st'' (ctx_sp ctx) (ctx_ge ctx)) (abstract_sequence empty b) st'. -Proof. - induction b using rev_ind; simplify. - { econstructor. simplify. rewrite app_nil_r in H. eauto. apply abstract_interp_empty. } - { simplify. rewrite app_assoc in H. rewrite abstract_seq in H. - exploit sem_update'; eauto; simplify. - exploit IHb; eauto; inv_simp. - econstructor; split; eauto. - rewrite abstract_seq. - eapply sem_update2; eauto. - } -Qed.*) - -Lemma sem_update_RBnop : - forall A ctx f st', - @sem A ctx f st' -> sem ctx (update f RBnop) st'. -Proof. auto. Qed. - -Lemma sem_update_Op : - forall A ge sp ist f st' r l o0 o m rs v ps, - @sem A (mk_ctx ist sp ge) f st' -> - eval_predf ps o = true -> - Op.eval_operation ge sp o0 (rs ## l) m = Some v -> - match_states st' (mk_instr_state rs ps m) -> - exists tst, - sem (mk_ctx ist sp ge) (update f (RBop (Some o) o0 l r)) tst - /\ match_states (mk_instr_state (Regmap.set r v rs) ps m) tst. -Proof. - intros. inv H1. inv H. inv H1. inv H3. simplify. - econstructor. simplify. - { constructor; try constructor; intros; try solve [rewrite genmap1; now eauto]. - destruct (Pos.eq_dec x r); subst. - { rewrite map2. specialize (H1 r). inv H1. -(*} - } - destruct st. - econstructor. simplify. - { constructor. - { constructor. intros. destruct (Pos.eq_dec x r); subst. - { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. - { inv H9. eapply gen_list_base; eauto. } - { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } - { inv H. rewrite genmap1; crush. eauto. } } - { constructor; eauto. intros. - destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed.*) Admitted. - -Lemma sem_update : - forall A ge sp st x st' st'' st''' f, - sem (mk_ctx st sp ge) f st' -> - 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. - 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', - RTLBlock.step_instr_list ge sp st bb st' -> - forall tst, - match_states st tst -> - exists tst', sem (mk_ctx tst sp ge) (abstract_sequence empty bb) tst' - /\ match_states st' tst'. -Proof. - induction bb using rev_ind; simplify. - { econstructor. simplify. apply abstract_interp_empty. - inv H. auto. } - { apply rtlblock_trans_correct' in H. simplify. - rewrite abstract_seq. - exploit IHbb; try eassumption; []; simplify. - exploit sem_update. apply H1. symmetry; eassumption. - eauto. simplify. econstructor. split. apply H3. - auto. } -Qed. - -Lemma abstract_execution_correct: - forall bb bb' cfi cfi' ge tge sp st st' tst, - RTLBlock.step_instr_list ge sp st bb st' -> - ge_preserved ge tge -> - schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi') = true -> - match_states st tst -> - exists tst', RTLPar.step_instr_block tge sp tst bb' tst' - /\ match_states st' tst'. -Proof. - intros. - unfold schedule_oracle in *. simplify. unfold empty_trees in H4. - exploit rtlblock_trans_correct; try eassumption; []; simplify. -(*) exploit abstract_execution_correct'; - try solve [eassumption | apply state_lessdef_match_sem; eassumption]. - apply match_states_commut. eauto. inv_simp. - exploit rtlpar_trans_correct; try eassumption; []; inv_simp. - exploit step_instr_block_matches; eauto. apply match_states_commut; eauto. inv_simp. - repeat match goal with | H: match_states _ _ |- _ => inv H end. - do 2 econstructor; eauto. - econstructor; congruence. -Qed.*)Admitted. - -Definition match_prog (prog : RTLBlock.program) (tprog : RTLPar.program) := - match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq prog tprog. - -Inductive match_stackframes: RTLBlock.stackframe -> RTLPar.stackframe -> Prop := -| match_stackframe: - forall f tf res sp pc rs rs' ps ps', - transl_function f = OK tf -> - (forall x, rs !! x = rs' !! x) -> - (forall x, ps !! x = ps' !! x) -> - match_stackframes (Stackframe res f sp pc rs ps) - (Stackframe res tf sp pc rs' ps'). - -Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop := -| match_state: - forall sf f sp pc rs rs' m sf' tf ps ps' - (TRANSL: transl_function f = OK tf) - (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_returnstate: - forall stack stack' v m - (STACKS: list_forall2 match_stackframes stack stack'), - match_states (Returnstate stack v m) - (Returnstate stack' v m) -| match_callstate: - forall stack stack' f tf args m - (TRANSL: transl_fundef f = OK tf) - (STACKS: list_forall2 match_stackframes stack stack'), - match_states (Callstate stack f args m) - (Callstate stack' tf args m). - -Section CORRECTNESS. - - Context (prog: RTLBlock.program) (tprog : RTLPar.program). - Context (TRANSL: match_prog prog tprog). - - Let ge : RTLBlock.genv := Globalenvs.Genv.globalenv prog. - Let tge : RTLPar.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. - Hint Resolve symbols_preserved : rtlgp. - - Lemma function_ptr_translated: - forall (b: Values.block) (f: RTLBlock.fundef), - Genv.find_funct_ptr ge b = Some f -> - exists tf, - Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = Errors.OK tf. - Proof using TRANSL. - intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. - intros (cu & tf & P & Q & R); exists tf; auto. - Qed. - - Lemma functions_translated: - forall (v: Values.val) (f: RTLBlock.fundef), - Genv.find_funct ge v = Some f -> - exists tf, - Genv.find_funct tge v = Some tf /\ transl_fundef f = Errors.OK tf. - Proof using TRANSL. - intros. exploit (Genv.find_funct_match TRANSL); eauto. - intros (cu & tf & P & Q & R); exists tf; auto. - Qed. - - Lemma senv_preserved: - Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). - Proof (Genv.senv_transf_partial TRANSL). - Hint Resolve senv_preserved : rtlgp. - - Lemma sig_transl_function: - forall (f: RTLBlock.fundef) (tf: RTLPar.fundef), - transl_fundef f = OK tf -> - funsig tf = funsig f. - Proof using . - unfold transl_fundef, transf_partial_fundef, transl_function; intros; - repeat destruct_match; crush; - match goal with H: OK _ = OK _ |- _ => inv H end; auto. - Qed. - Hint Resolve sig_transl_function : rtlgp. - - Hint Resolve Val.lessdef_same : rtlgp. - Hint Resolve regs_lessdef_regs : rtlgp. - - Lemma find_function_translated: - forall ros rs rs' f, - (forall x, rs !! x = rs' !! x) -> - find_function ge ros rs = Some f -> - exists tf, find_function tge ros rs' = Some tf - /\ transl_fundef f = OK tf. - Proof using TRANSL. - Ltac ffts := match goal with - | [ H: forall _, Val.lessdef _ _, r: Registers.reg |- _ ] => - specialize (H r); inv H - | [ H: Vundef = ?r, H1: Genv.find_funct _ ?r = Some _ |- _ ] => - rewrite <- H in H1 - | [ H: Genv.find_funct _ Vundef = Some _ |- _] => solve [inv H] - | _ => solve [exploit functions_translated; eauto] - end. - destruct ros; simplify; try rewrite <- H; - [| rewrite symbols_preserved; destruct_match; - try (apply function_ptr_translated); crush ]; - intros; - repeat ffts. - Qed. - - Lemma schedule_oracle_nil: - forall bb cfi, - schedule_oracle {| bb_body := nil; bb_exit := cfi |} bb = true -> - bb_body bb = nil /\ bb_exit bb = cfi. - Proof using . - unfold schedule_oracle, check_control_flow_instr. - simplify; repeat destruct_match; crush. - Qed. - - Lemma schedule_oracle_nil2: - forall cfi, - schedule_oracle {| bb_body := nil; bb_exit := cfi |} - {| bb_body := nil; bb_exit := cfi |} = true. - Proof using . - unfold schedule_oracle, check_control_flow_instr. - simplify; repeat destruct_match; crush. - Qed. - - Lemma eval_op_eq: - forall (sp0 : Values.val) (op : Op.operation) (vl : list Values.val) m, - Op.eval_operation ge sp0 op vl m = Op.eval_operation tge sp0 op vl m. - Proof using TRANSL. - intros. - destruct op; auto; unfold Op.eval_operation, Genv.symbol_address, Op.eval_addressing32; - [| destruct a; unfold Genv.symbol_address ]; - try rewrite symbols_preserved; auto. - Qed. - Hint Resolve eval_op_eq : rtlgp. - - Lemma eval_addressing_eq: - forall sp addr vl, - Op.eval_addressing ge sp addr vl = Op.eval_addressing tge sp addr vl. - Proof using TRANSL. - intros. - destruct addr; - unfold Op.eval_addressing, Op.eval_addressing32; - unfold Genv.symbol_address; - try rewrite symbols_preserved; auto. - Qed. - Hint Resolve eval_addressing_eq : rtlgp. - - Lemma ge_preserved_lem: - ge_preserved ge tge. - Proof using TRANSL. - unfold ge_preserved. - eauto with rtlgp. - Qed. - Hint Resolve ge_preserved_lem : rtlgp. - - Lemma lessdef_regmap_optget: - forall or rs rs', - regs_lessdef rs rs' -> - Val.lessdef (regmap_optget or Vundef rs) (regmap_optget or Vundef rs'). - Proof using. destruct or; crush. Qed. - Hint Resolve lessdef_regmap_optget : rtlgp. - - Lemma regmap_equiv_lessdef: - forall rs rs', - (forall x, rs !! x = rs' !! x) -> - regs_lessdef rs rs'. - Proof using. - intros; unfold regs_lessdef; intros. - rewrite H. apply Val.lessdef_refl. - Qed. - Hint Resolve regmap_equiv_lessdef : rtlgp. - - Lemma int_lessdef: - forall rs rs', - regs_lessdef rs rs' -> - (forall arg v, - rs !! arg = Vint v -> - rs' !! arg = Vint v). - Proof using. intros ? ? H; intros; specialize (H arg); inv H; crush. Qed. - Hint Resolve int_lessdef : rtlgp. - - Ltac semantics_simpl := - match goal with - | [ H: match_states _ _ |- _ ] => - let H2 := fresh "H" in - learn H as H2; inv H2 - | [ H: transl_function ?f = OK _ |- _ ] => - let H2 := fresh "TRANSL" in - learn H as H2; - unfold transl_function in H2; - destruct (check_scheduled_trees - (@fn_code RTLBlock.bb f) - (@fn_code RTLPar.bb (schedule f))) eqn:?; - [| discriminate ]; inv H2 - | [ H: context[check_scheduled_trees] |- _ ] => - let H2 := fresh "CHECK" in - learn H as H2; - eapply check_scheduled_trees_correct in H2; [| solve [eauto] ] - | [ H: schedule_oracle {| bb_body := nil; bb_exit := _ |} ?bb = true |- _ ] => - let H2 := fresh "SCHED" in - learn H as H2; - apply schedule_oracle_nil in H2 - | [ H: find_function _ _ _ = Some _, H2: forall x, ?rs !! x = ?rs' !! x |- _ ] => - learn H; exploit find_function_translated; try apply H2; eauto; inversion 1 - | [ H: Mem.free ?m _ _ _ = Some ?m', H2: Mem.extends ?m ?m'' |- _ ] => - learn H; exploit Mem.free_parallel_extends; eauto; intros - | [ H: Events.eval_builtin_args _ _ _ _ _ _, H2: regs_lessdef ?rs ?rs' |- _ ] => - let H3 := fresh "H" in - learn H; exploit Events.eval_builtin_args_lessdef; [apply H2 | | |]; - eauto with rtlgp; intro H3; learn H3 - | [ H: Events.external_call _ _ _ _ _ _ _ |- _ ] => - let H2 := fresh "H" in - learn H; exploit Events.external_call_mem_extends; - eauto; intro H2; learn H2 - | [ H: exists _, _ |- _ ] => inv H - | _ => progress simplify - end. - - Hint Resolve Events.eval_builtin_args_preserved : rtlgp. - Hint Resolve Events.external_call_symbols_preserved : rtlgp. - Hint Resolve set_res_lessdef : rtlgp. - Hint Resolve set_reg_lessdef : rtlgp. - Hint Resolve Op.eval_condition_lessdef : rtlgp. - - Hint Constructors Events.eval_builtin_arg: barg. - - Lemma eval_builtin_arg_eq: - forall A ge a v1 m1 e1 e2 sp, - (forall x, e1 x = e2 x) -> - @Events.eval_builtin_arg A ge e1 sp m1 a v1 -> - Events.eval_builtin_arg ge e2 sp m1 a v1. -Proof. induction 2; try rewrite H; eauto with barg. Qed. - - Lemma eval_builtin_args_eq: - forall A ge e1 sp m1 e2 al vl1, - (forall x, e1 x = e2 x) -> - @Events.eval_builtin_args A ge e1 sp m1 al vl1 -> - Events.eval_builtin_args ge e2 sp m1 al vl1. - Proof. - induction 2. - - econstructor; split. - - exploit eval_builtin_arg_eq; eauto. intros. - destruct IHlist_forall2 as [| y]. constructor; eauto. - constructor. constructor; auto. - constructor; eauto. - Qed. - - Lemma step_cf_instr_correct: - forall cfi t s s', - step_cf_instr ge s cfi t s' -> - forall r, - match_states s r -> - exists r', step_cf_instr tge r cfi t r' /\ match_states s' r'. - Proof using TRANSL. - induction 1; repeat semantics_simpl; - try solve [repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp)]. - { do 3 (try erewrite match_states_list by eauto; econstructor; eauto with rtlgp). - eapply eval_builtin_args_eq. eapply REG. - eapply Events.eval_builtin_args_preserved. eapply symbols_preserved. - eauto. - intros. - unfold regmap_setres. destruct res. - destruct (Pos.eq_dec x0 x); subst. - repeat rewrite Regmap.gss; auto. - repeat rewrite Regmap.gso; auto. - eapply REG. eapply REG. - } - { repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). - unfold regmap_optget. destruct or. rewrite REG. constructor; eauto. - constructor; eauto. - } - { exploit IHstep_cf_instr; eauto. simplify. - repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). - erewrite eval_predf_pr_equiv; eauto. - } - Qed. - - Theorem transl_step_correct : - forall (S1 : RTLBlock.state) t S2, - RTLBlock.step ge S1 t S2 -> - forall (R1 : RTLPar.state), - match_states S1 R1 -> - exists R2, Smallstep.plus RTLPar.step tge R1 t R2 /\ match_states S2 R2. - Proof. - - induction 1; repeat semantics_simpl. - - { destruct bb; destruct x. - assert (bb_exit = bb_exit0). - { unfold schedule_oracle in *. simplify. - unfold check_control_flow_instr in *. - destruct_match; crush. - } - subst. - - exploit abstract_execution_correct; try eassumption. eapply ge_preserved_lem. - econstructor; eauto. - simplify. destruct x. inv H7. - - exploit step_cf_instr_correct; try eassumption. econstructor; eauto. - simplify. - - econstructor. econstructor. eapply Smallstep.plus_one. econstructor. - 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. - 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. } - { inv STACKS. inv H2. repeat econstructor; eauto. - intros. apply PTree_matches; eauto. } - Qed. - - Lemma transl_initial_states: - forall S, - RTLBlock.initial_state prog S -> - exists R, RTLPar.initial_state tprog R /\ match_states S R. - Proof. - induction 1. - exploit function_ptr_translated; eauto. intros [tf [A B]]. - econstructor; split. - econstructor. apply (Genv.init_mem_transf_partial TRANSL); eauto. - replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto. - symmetry; eapply match_program_main; eauto. - eexact A. - rewrite <- H2. apply sig_transl_function; auto. - constructor. auto. constructor. - Qed. - - Lemma transl_final_states: - forall S R r, - match_states S R -> RTLBlock.final_state S r -> RTLPar.final_state R r. - Proof. - intros. inv H0. inv H. inv STACKS. constructor. - Qed. - - Theorem transf_program_correct: - Smallstep.forward_simulation (RTLBlock.semantics prog) (RTLPar.semantics tprog). - Proof. - eapply Smallstep.forward_simulation_plus. - apply senv_preserved. - eexact transl_initial_states. - eexact transl_final_states. - exact transl_step_correct. - Qed. - -End CORRECTNESS. -#+end_src - -* License - -#+name: license -#+begin_src coq :tangle no -(* - * 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 . - *) -#+end_src 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 18640c782f216e6e62e25ce24b6061ad93703cf6 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 24 Mar 2022 11:03:10 +0000 Subject: Update makefile and remove _CoqProject --- Makefile | 19 +++++++++++-------- _CoqProject | 15 --------------- 2 files changed, 11 insertions(+), 23 deletions(-) delete mode 100644 _CoqProject diff --git a/Makefile b/Makefile index 3b51d00..978a109 100644 --- a/Makefile +++ b/Makefile @@ -22,11 +22,11 @@ COQMAKE := $(COQBIN)coq_makefile COQDOCFLAGS := --no-lib-name -l VS := src/Compiler.v src/Simulator.v src/HLSOpts.v $(foreach d, common hls bourdoncle, $(wildcard src/$(d)/*.v)) -LIT := $(wildcard lit/*.org) +LIT := docs/basic-block-generation.org docs/scheduler.org docs/scheduler-languages.org PREFIX ?= . -.PHONY: all install proof clean extraction test +.PHONY: all install proof clean extraction test force all: lib/COMPCERTSTAMP $(MAKE) proof @@ -74,24 +74,27 @@ src/extraction/STAMP: @$(COQEXEC) ./src/extraction/Extraction.v @touch $@ -Makefile.coq: +Makefile.coq: force @echo "COQMAKE Makefile.coq" $(COQBIN)coq_makefile $(COQINCLUDES) $(VS) -o Makefile.coq + echo "$(COQINCLUDES)" >_CoqProject + +force: docs/vericert.1: $(MAKE) -C docs vericert.1 -clean:: Makefile.coq - $(MAKE) -f Makefile.coq clean - $(MAKE) -C test clean - rm -f Makefile.coq - detangle-all: emacs --batch --eval "(progn(require 'org)(require 'ob-tangle)\ (setq org-src-preserve-indentation t)\ $(foreach vs,$(VS),(org-babel-detangle \"$(vs)\"))\ (org-save-all-org-buffers))" +clean:: Makefile.coq + $(MAKE) -f Makefile.coq clean + $(MAKE) -C test clean + rm -f Makefile.coq + clean:: rm -f */*.v.d */*.glob */*.vo */*~ *~ rm -f src/extraction/*.ml src/extraction/*.mli diff --git a/_CoqProject b/_CoqProject deleted file mode 100644 index 69d12c1..0000000 --- a/_CoqProject +++ /dev/null @@ -1,15 +0,0 @@ --R src/common vericert.common --R src/extraction vericert.extraction --R src/hls vericert.hls --R src vericert - --R lib/CompCert/backend compcert.backend --R lib/CompCert/cfrontend compcert.cfrontend --R lib/CompCert/common compcert.common --R lib/CompCert/cparser compcert.cparser --R lib/CompCert/driver compcert.driver --R lib/CompCert/exportclight compcert.exportclight --R lib/CompCert/lib compcert.lib --R lib/CompCert/verilog compcert.verilog --R lib/CompCert/flocq Flocq --R lib/CompCert/MenhirLib MenhirLib -- cgit From b9e793eef135bd411c9945f0b1ba99308d9edbd5 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 24 Mar 2022 11:45:42 +0000 Subject: Force _CoqProject generation and fix namespaces --- Makefile | 7 ++----- 1 file changed, 2 insertions(+), 5 deletions(-) diff --git a/Makefile b/Makefile index 978a109..de3c3b7 100644 --- a/Makefile +++ b/Makefile @@ -8,10 +8,7 @@ endif COMPCERTRECDIRS := lib common verilog backend cfrontend driver cparser -COQINCLUDES := -R src/common vericert.common \ - -R src/extraction vericert.extraction \ - -R src/hls vericert.hls \ - -R src vericert \ +COQINCLUDES := -R src vericert \ $(foreach d, $(COMPCERTRECDIRS), -R lib/CompCert/$(d) compcert.$(d)) \ -R lib/CompCert/flocq Flocq \ -R lib/CompCert/MenhirLib MenhirLib @@ -74,7 +71,7 @@ src/extraction/STAMP: @$(COQEXEC) ./src/extraction/Extraction.v @touch $@ -Makefile.coq: force +Makefile.coq _CoqProject: force @echo "COQMAKE Makefile.coq" $(COQBIN)coq_makefile $(COQINCLUDES) $(VS) -o Makefile.coq echo "$(COQINCLUDES)" >_CoqProject -- 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 --- docs/scheduler-languages.org | 199 +++++++++++++++++++++---------------------- docs/scheduler.org | 54 ++++++++---- 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 +- 7 files changed, 207 insertions(+), 268 deletions(-) diff --git a/docs/scheduler-languages.org b/docs/scheduler-languages.org index cf03b3a..02249f8 100644 --- a/docs/scheduler-languages.org +++ b/docs/scheduler-languages.org @@ -224,8 +224,18 @@ Section DEFINITION. (rs: regset) (**r register state in calling function *) (pr: predset), (**r predicate state of the calling function *) stackframe. +#+end_src + +*** State definition +:PROPERTIES: +:CUSTOM_ID: state +:END: - Inductive state : Type := +Definition of the ~state~ type, which is used by the ~step~ functions. + +#+name: rtlblockinstr-state +#+begin_src coq + Variant state : Type := | State: forall (stack: list stackframe) (**r call stack *) (f: function) (**r current function *) @@ -246,7 +256,10 @@ Section DEFINITION. (v: val) (**r return value for the call *) (m: mem), (**r memory state *) state. +#+end_src +#+name: rtlblockinstr-state +#+begin_src coq End DEFINITION. #+end_src @@ -284,7 +297,15 @@ Section RELSEM. eval_pred (Some p) i i' i | eval_pred_none: forall i i', eval_pred None i i' i. +#+end_src + +*** Step a single instruction +:PROPERTIES: +:CUSTOM_ID: step_instr +:END: +#+name: rtlblockinstr-step_instr +#+begin_src coq Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := | exec_RBnop: forall sp ist, @@ -311,7 +332,15 @@ 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. +#+end_src + +*** Step a control-flow instruction +:PROPERTIES: +:CUSTOM_ID: step_cf_instr +:END: +#+name: rtlblockinstr-step_cf_instr +#+begin_src coq 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, @@ -356,7 +385,10 @@ 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'. +#+end_src +#+name: rtlblockinstr-end_RELSEM +#+begin_src coq End RELSEM. #+end_src @@ -394,137 +426,98 @@ Definition fundef := @fundef bb. Definition program := @program bb. Definition funsig := @funsig bb. Definition stackframe := @stackframe bb. +Definition state := @state bb. Definition genv := @genv bb. +#+end_src -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. +** Semantics + +We first describe the semantics by assuming a global program environment with type ~genv~ which was +declared earlier. +#+name: rtlblock-semantics +#+begin_src coq Section RELSEM. Context (ge: genv). +#+end_src + +*** Instruction list step +:PROPERTIES: +:CUSTOM_ID: step_instr_list +:END: + +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~ ([[#step_instr][step_instr]]). +#+name: rtlblock-step_instr_list +#+begin_src coq 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. + | 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. +#+end_src - 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. +*** Top-level step +:PROPERTIES: +:CUSTOM_ID: rtlblock-step +: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'. +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. - Inductive step: state -> trace -> state -> Prop := +#+name: rtlblock-step +#+begin_src coq + 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). +#+end_src +#+name: rtlblock-rest +#+begin_src coq 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). diff --git a/docs/scheduler.org b/docs/scheduler.org index 018633c..c8b00ff 100644 --- a/docs/scheduler.org +++ b/docs/scheduler.org @@ -909,18 +909,19 @@ 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. +#+end_src -(** ** 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. +#+name: rtlpargen-generation +#+begin_src coq Fixpoint list_translation (l : list reg) (f : forest) {struct l} : list pred_expr := match l with | nil => nil @@ -1014,18 +1015,25 @@ 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 +#+end_src -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: +*** Update Function +:PROPERTIES: +:CUSTOM_ID: update-function +:END: + +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. *) +the other predicates. +#+name: rtlpargen-update-function +#+begin_src coq Definition update (f : forest) (i : instr) : forest := match i with | RBnop => f @@ -1055,28 +1063,37 @@ Definition update (f : forest) (i : instr) : forest := (f # (Pred p)) (map_predicated (pred_ret (Esetpred c)) (merge (list_translation args f)))) end. +#+end_src -(** Implementing which are necessary to show the correctness of the translation validation by +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. +#+name: rtlpargen-abstract-seq +#+begin_src coq Fixpoint abstract_sequence (f : forest) (b : list instr) : forest := match b with | nil => f | i :: l => abstract_sequence (update f i) l end. +#+end_src -(** Check equivalence of control flow instructions. As none of the basic blocks should have been +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. *) +exactly. +#+name: rtlpargen-check-functions +#+begin_src coq Definition check_control_flow_instr (c1 c2: cf_instr) : bool := if cf_instr_eq c1 c2 then true else false. +#+end_src -(** 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. +#+name: rtlpargen-top-check-functions +#+begin_src coq Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := match bb with | nil => @@ -1116,8 +1133,12 @@ Lemma check_scheduled_trees_correct2: PTree.get x f2 = None. Proof. solve_scheduled_trees_correct. Qed. -(** ** Top-level Functions *) +#+end_src + +** Top-level Functions +#+name: rtlpargen-top-level-functions +#+begin_src coq Parameter schedule : RTLBlock.function -> RTLPar.function. Definition transl_function (f: RTLBlock.function) : Errors.res RTLPar.function := @@ -1136,6 +1157,7 @@ Definition transl_fundef := transf_partial_fundef transl_function. Definition transl_program (p : RTLBlock.program) : Errors.res RTLPar.program := transform_partial_program transl_fundef p. #+end_src + * RTLPargenproof :PROPERTIES: :header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargenproof.v 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 --- Makefile | 9 +- default.nix | 4 + docs/basic-block-generation.org | 497 --------- docs/scheduler-languages.org | 682 ------------ docs/scheduler.org | 2317 --------------------------------------- src/Compiler.v | 115 +- src/bourdoncle/Bourdoncle.v | 4 +- src/common/Maps.v | 5 +- 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 +- 26 files changed, 582 insertions(+), 3770 deletions(-) delete mode 100644 docs/basic-block-generation.org delete mode 100644 docs/scheduler-languages.org delete mode 100644 docs/scheduler.org diff --git a/Makefile b/Makefile index de3c3b7..a6aa914 100644 --- a/Makefile +++ b/Makefile @@ -17,6 +17,7 @@ COQEXEC := $(COQBIN)coqtop $(COQINCLUDES) -batch -load-vernac-source COQMAKE := $(COQBIN)coq_makefile COQDOCFLAGS := --no-lib-name -l +ALECTRYON_OPTS := --html-minification --long-line-threshold 80 --coq-driver sertop_noexec $(COQINCLUDES) VS := src/Compiler.v src/Simulator.v src/HLSOpts.v $(foreach d, common hls bourdoncle, $(wildcard src/$(d)/*.v)) LIT := docs/basic-block-generation.org docs/scheduler.org docs/scheduler-languages.org @@ -48,7 +49,13 @@ proof: Makefile.coq $(MAKE) -f Makefile.coq @rm -f src/extraction/STAMP -doc: Makefile.coq +doc-html: $(patsubst src/%.v,html/%.html,$(VS)) + +html/%.html: src/%.v + @mkdir -p $(dir $@) + @echo "ALECTRYON" $@; alectryon $(ALECTRYON_OPTS) $< -o $@ + +coqdoc: Makefile.coq $(MAKE) COQDOCFLAGS="$(COQDOCFLAGS)" -f Makefile.coq html cp ./docs/res/coqdoc.css html/. diff --git a/default.nix b/default.nix index e3ce09b..eb1125c 100644 --- a/default.nix +++ b/default.nix @@ -11,6 +11,10 @@ stdenv.mkDerivation { ncoq.ocaml ncoq.ocamlPackages.findlib ncoq.ocamlPackages.menhir ncoq.ocamlPackages.ocamlgraph ncoq.ocamlPackages.merlin ncoq.ocamlPackages.menhirLib + + ncoqPackages.serapi + python3 + python3Packages.alectryon ]; enableParallelBuilding = true; diff --git a/docs/basic-block-generation.org b/docs/basic-block-generation.org deleted file mode 100644 index 1d78dad..0000000 --- a/docs/basic-block-generation.org +++ /dev/null @@ -1,497 +0,0 @@ -#+title: Basic Block Generation -#+author: Yann Herklotz -#+email: yann [at] yannherklotz [dot] com - -* RTLBlockgen -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgen.v -:END: - -Refers to [[rtlblockgen-equalities][rtlblockgen-equalities]]. - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblockgen-imports -#+begin_src coq -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. -#+end_src - -#+name: rtlblockgen-equalities-insert -#+begin_src coq :comments no -<> -#+end_src - -#+name: rtlblockgen-main -#+begin_src coq -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 - | 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. -#+end_src - -** Equalities - -#+name: rtlblockgen-equalities -#+begin_src coq :tangle no -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. -#+end_src - -* RTLBlockgenproof -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockgenproof.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -** Imports - -#+name: rtlblockgenproof-imports -#+begin_src coq -Require compcert.backend.RTL. -Require Import compcert.common.AST. -Require Import compcert.lib.Maps. - -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLBlockgen. -#+end_src - -** Match states - -The ~match_states~ predicate describes which states are equivalent between the two languages, in this -case ~RTL~ and ~RTLBlock~. - -#+name: rtlblockgenproof-match-states -#+begin_src coq -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). -#+end_src - -** Correctness - -#+name: rtlblockgenproof-correctness -#+begin_src coq -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. -#+end_src - -* Partition -:PROPERTIES: -:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Partition.ml -:END: - -#+begin_src ocaml :comments no :padline no :exports none -<> -#+end_src - -#+name: partition-main -#+begin_src ocaml -open Printf -open Clflags -open Camlcoq -open Datatypes -open Coqlib -open Maps -open AST -open Kildall -open Op -open RTLBlockInstr -open RTLBlock - -(** 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 - let filt = List.filter (fun s -> P.lt n s || P.lt s (P.pred n)) succ in - ((match filt with [] -> [] | _ -> [n]), filt) - -let find_edges c = - PTree.fold (fun l n i -> - let f = find_edge i n in - (List.append (fst f) (fst l), List.append (snd f) (snd l))) c ([], []) - -let prepend_instr i = function - | {bb_body = bb; bb_exit = e} -> {bb_body = (i :: bb); bb_exit = e} - -let translate_inst = function - | RTL.Inop _ -> Some RBnop - | RTL.Iop (op, ls, dst, _) -> Some (RBop (None, op, ls, dst)) - | RTL.Iload (m, addr, ls, dst, _) -> Some (RBload (None, m, addr, ls, dst)) - | RTL.Istore (m, addr, ls, src, _) -> Some (RBstore (None, m, addr, ls, src)) - | _ -> None - -let translate_cfi = function - | RTL.Icall (s, r, ls, dst, n) -> Some (RBcall (s, r, ls, dst, n)) - | RTL.Itailcall (s, r, ls) -> Some (RBtailcall (s, r, ls)) - | RTL.Ibuiltin (e, ls, r, n) -> Some (RBbuiltin (e, ls, r, n)) - | RTL.Icond (c, ls, dst1, dst2) -> Some (RBcond (c, ls, dst1, dst2)) - | RTL.Ijumptable (r, ls) -> Some (RBjumptable (r, ls)) - | RTL.Ireturn r -> Some (RBreturn r) - | _ -> None - -let rec next_bblock_from_RTL is_start e (c : RTL.code) s i = - let succ = List.map (fun i -> (i, PTree.get i c)) (RTL.successors_instr i) in - let trans_inst = (translate_inst i, translate_cfi i) in - 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 = [RBnop]; bb_exit = RBgoto s } - else - 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 = [RBnop]; bb_exit = RBgoto s } - else begin - match next_bblock_from_RTL false e c s' i_n with - | Errors.OK bb -> - Errors.OK (prepend_instr i' bb) - | Errors.Error msg -> Errors.Error msg - end - | _, _ -> - Errors.Error (Errors.msg (coqstring_of_camlstring "next_bblock_from_RTL went wrong.")) - -let rec traverseacc f l c = - match l with - | [] -> Errors.OK c - | x::xs -> - match f x c with - | Errors.Error msg -> Errors.Error msg - | Errors.OK x' -> - match traverseacc f xs x' with - | Errors.Error msg -> Errors.Error msg - | Errors.OK xs' -> Errors.OK xs' - -let rec translate_all edge c s res = - let c_bb, translated = res in - if List.exists (fun x -> P.eq x s) translated then Errors.OK (c_bb, translated) else - (match PTree.get s c with - | None -> Errors.Error (Errors.msg (coqstring_of_camlstring "Could not translate all.")) - | Some i -> - match next_bblock_from_RTL true edge c s i with - | Errors.Error msg -> Errors.Error msg - | Errors.OK {bb_body = bb; bb_exit = e} -> - let succ = List.filter (fun x -> P.lt x s) (successors_instr e) in - (match traverseacc (translate_all edge c) succ (c_bb, s :: translated) with - | Errors.Error msg -> Errors.Error msg - | Errors.OK (c', t') -> - Errors.OK (PTree.set s {bb_body = bb; bb_exit = e} c', t'))) - -(* Partition a function and transform it into RTLBlock. *) -let function_from_RTL f = - let e = find_edges f.RTL.fn_code in - match translate_all e f.RTL.fn_code f.RTL.fn_entrypoint (PTree.empty, []) with - | Errors.Error msg -> Errors.Error msg - | Errors.OK (c, _) -> - Errors.OK { fn_sig = f.RTL.fn_sig; - fn_stacksize = f.RTL.fn_stacksize; - fn_params = f.RTL.fn_params; - fn_entrypoint = f.RTL.fn_entrypoint; - fn_code = c - } - -let partition = function_from_RTL -#+end_src - -* License - -#+name: license -#+begin_src coq :tangle no -(* - * 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 . - *) -#+end_src diff --git a/docs/scheduler-languages.org b/docs/scheduler-languages.org deleted file mode 100644 index 02249f8..0000000 --- a/docs/scheduler-languages.org +++ /dev/null @@ -1,682 +0,0 @@ -#+title: Scheduler Languages -#+author: Yann Herklotz -#+email: yann [at] yannherklotz [dot] com - -* RTLBlockInstr -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlockInstr.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblockinstr-imports -#+begin_src coq -Require Import Coq.micromega.Lia. - -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.Values. -Require Import compcert.lib.Integers. -Require Import compcert.lib.Maps. -Require Import compcert.verilog.Op. - -Require Import Predicate. -Require Import Vericertlib. -#+end_src - -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]). - -#+name: rtlblockinstr-instr-def -#+begin_src coq -Definition node := positive. - -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. -#+end_src - -** Control-Flow Instruction Definition - -These are the instructions that count as control-flow, and will be placed at the end of the basic -blocks. - -#+name: rtlblockinstr-cf-instr-def -#+begin_src coq -Inductive cf_instr : Type := -| RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr -| RBtailcall : signature -> reg + ident -> list reg -> cf_instr -| RBbuiltin : external_function -> list (builtin_arg reg) -> - builtin_res reg -> node -> cf_instr -| RBcond : condition -> list reg -> node -> node -> cf_instr -| RBjumptable : reg -> list node -> cf_instr -| RBreturn : option reg -> cf_instr -| RBgoto : node -> cf_instr -| RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. -#+end_src - -** Helper functions - -#+name: rtlblockinstr-helpers -#+begin_src coq -Fixpoint successors_instr (i : cf_instr) : list node := - match i with - | RBcall sig ros args res s => s :: nil - | RBtailcall sig ros args => nil - | RBbuiltin ef args res s => s :: nil - | RBcond cond args ifso ifnot => ifso :: ifnot :: nil - | RBjumptable arg tbl => tbl - | RBreturn optarg => nil - | RBgoto n => n :: nil - | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil) - end. - -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) - | RBload p chunk addr args dst => - fold_left Pos.max args (Pos.max dst m) - | RBstore p chunk addr args src => - fold_left Pos.max args (Pos.max src m) - | RBsetpred p' c args p => - 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)) - | RBcall sig (inr id) args res s => - fold_left Pos.max args (Pos.max res m) - | RBtailcall sig (inl r) args => - fold_left Pos.max args (Pos.max r m) - | RBtailcall sig (inr id) args => - 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) - | RBcond cond args ifso ifnot => fold_left Pos.max args m - | RBjumptable arg tbl => Pos.max arg m - | RBreturn None => m - | RBreturn (Some arg) => Pos.max arg m - | RBgoto n => m - | RBpred_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. -#+end_src - -*** 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. - -#+name: rtlblockinstr-instr-state -#+begin_src coq -Record instr_state := mk_instr_state { - is_rs: regset; - is_ps: predset; - is_mem: mem; -}. -#+end_src - -** Top-Level Type Definitions - -#+name: rtlblockinstr-type-def -#+begin_src coq -Section DEFINITION. - - Context {bblock_body: Type}. - - 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_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. -#+end_src - -*** State definition -:PROPERTIES: -:CUSTOM_ID: state -:END: - -Definition of the ~state~ type, which is used by the ~step~ functions. - -#+name: rtlblockinstr-state -#+begin_src coq - Variant 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. -#+end_src - -#+name: rtlblockinstr-state -#+begin_src coq -End DEFINITION. -#+end_src - -** Semantics - -#+name: rtlblockinstr-semantics -#+begin_src coq -Section RELSEM. - - Context {bblock_body : Type}. - - Definition genv := Genv.t (@fundef bblock_body) unit. - - 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. -#+end_src - -*** Step a single instruction -:PROPERTIES: -:CUSTOM_ID: step_instr -:END: - -#+name: rtlblockinstr-step_instr -#+begin_src coq - Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := - | exec_RBnop: - 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 - | 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 - | 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 - | 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. -#+end_src - -*** Step a control-flow instruction -:PROPERTIES: -:CUSTOM_ID: step_cf_instr -:END: - -#+name: rtlblockinstr-step_cf_instr -#+begin_src coq - 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, - 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') - 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) (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') - | 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) - | 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) - | 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) - 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) - | 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'. -#+end_src - -#+name: rtlblockinstr-end_RELSEM -#+begin_src coq -End RELSEM. -#+end_src - -* RTLBlock -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLBlock.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlblock-main -#+begin_src coq -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.RTLBlockInstr. - -Definition bb := list instr. - -Definition bblock := @bblock bb. -Definition code := @code bb. -Definition function := @function bb. -Definition fundef := @fundef bb. -Definition program := @program bb. -Definition funsig := @funsig bb. -Definition stackframe := @stackframe bb. -Definition state := @state bb. - -Definition genv := @genv bb. -#+end_src - -** Semantics - -We first describe the semantics by assuming a global program environment with type ~genv~ which was -declared earlier. - -#+name: rtlblock-semantics -#+begin_src coq -Section RELSEM. - - Context (ge: genv). -#+end_src - -*** Instruction list step -:PROPERTIES: -:CUSTOM_ID: step_instr_list -:END: - -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~ ([[#step_instr][step_instr]]). - -#+name: rtlblock-step_instr_list -#+begin_src coq - 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. -#+end_src - -*** Top-level step -:PROPERTIES: -:CUSTOM_ID: rtlblock-step -:END: - -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. - -#+name: rtlblock-step -#+begin_src coq - 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') - | 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_src - -#+name: rtlblock-rest -#+begin_src coq -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). -#+end_src - -* RTLPar -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPar.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpar-main -#+begin_src coq -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.RTLBlockInstr. - -Definition bb := list (list (list instr)). - -Definition bblock := @bblock bb. -Definition code := @code bb. -Definition function := @function bb. -Definition fundef := @fundef bb. -Definition program := @program bb. -Definition funsig := @funsig bb. -Definition stackframe := @stackframe bb. -Definition state := @state bb. -Definition genv := @genv bb. - -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_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 -> bb -> 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: 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 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') - | 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. -#+end_src - -* License - -#+name: license -#+begin_src coq :tangle no -(* - * 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 . - *) -#+end_src diff --git a/docs/scheduler.org b/docs/scheduler.org deleted file mode 100644 index c8b00ff..0000000 --- a/docs/scheduler.org +++ /dev/null @@ -1,2317 +0,0 @@ -#+title: Basic Block Generation -#+author: Yann Herklotz -#+email: yann [at] yannherklotz [dot] com - -* Scheduler -:PROPERTIES: -:header-args:ocaml: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/Schedule.ml -:END: - -#+begin_src ocaml :comments no :padline no :exports none -<> -#+end_src - -#+name: scheduler-main -#+begin_src ocaml -open Printf -open Clflags -open Camlcoq -open Datatypes -open Coqlib -open Maps -open AST -open Kildall -open Op -open RTLBlockInstr -open Predicate -open RTLBlock -open HTL -open Verilog -open HTLgen -open HTLMonad -open HTLMonadExtra - -module SS = Set.Make(P) - -type svtype = - | BBType of int - | VarType of int * int - -type sv = { - sv_type: svtype; - sv_num: int; -} - -let print_sv v = - match v with - | { sv_type = BBType bbi; sv_num = n } -> sprintf "bb%d_%d" bbi n - | { sv_type = VarType (bbi, i); sv_num = n } -> sprintf "var%dn%d_%d" bbi i n - -module G = Graph.Persistent.Digraph.ConcreteLabeled(struct - type t = sv - let compare = compare - let equal = (=) - let hash = Hashtbl.hash -end)(struct - type t = int - let compare = compare - let hash = Hashtbl.hash - let equal = (=) - let default = 0 -end) - -module GDot = Graph.Graphviz.Dot(struct - let graph_attributes _ = [] - let default_vertex_attributes _ = [] - let vertex_name = print_sv - let vertex_attributes _ = [] - let get_subgraph _ = None - let default_edge_attributes _ = [] - let edge_attributes _ = [] - - include G - end) - -module DFG = Graph.Persistent.Digraph.ConcreteLabeled(struct - type t = int * instr - let compare = compare - let equal = (=) - let hash = Hashtbl.hash -end)(struct - type t = int - let compare = compare - let hash = Hashtbl.hash - let equal = (=) - let default = 0 -end) - -module DFGSimp = Graph.Persistent.Graph.Concrete(struct - type t = int * instr - let compare = compare - let equal = (=) - let hash = Hashtbl.hash - end) - -let convert dfg = - DFG.fold_vertex (fun v g -> DFGSimp.add_vertex g v) dfg DFGSimp.empty - |> DFG.fold_edges (fun v1 v2 g -> DFGSimp.add_edge (DFGSimp.add_edge g v1 v2) v2 v1) dfg - -let reg r = sprintf "r%d" (P.to_int r) -let print_pred r = sprintf "p%d" (P.to_int r) - -let print_instr = function - | RBnop -> "" - | RBload (_, _, _, _, r) -> sprintf "load(%s)" (reg r) - | RBstore (_, _, _, _, r) -> sprintf "store(%s)" (reg r) - | RBsetpred (_, _, _, p) -> sprintf "setpred(%s)" (print_pred p) - | RBop (_, op, args, d) -> - (match op, args with - | Omove, _ -> "mov" - | Ointconst n, _ -> sprintf "%s=%ld" (reg d) (camlint_of_coqint n) - | Olongconst n, _ -> sprintf "%s=%LdL" (reg d) (camlint64_of_coqint n) - | Ofloatconst n, _ -> sprintf "%s=%.15F" (reg d) (camlfloat_of_coqfloat n) - | Osingleconst n, _ -> sprintf "%s=%.15Ff" (reg d) (camlfloat_of_coqfloat32 n) - | Oindirectsymbol id, _ -> sprintf "%s=&%s" (reg d) (extern_atom id) - | Ocast8signed, [r1] -> sprintf "%s=int8signed(%s)" (reg d) (reg r1) - | Ocast8unsigned, [r1] -> sprintf "%s=int8unsigned(%s)" (reg d) (reg r1) - | Ocast16signed, [r1] -> sprintf "%s=int16signed(%s)" (reg d) (reg r1) - | Ocast16unsigned, [r1] -> sprintf "%s=int16unsigned(%s)" (reg d) (reg r1) - | Oneg, [r1] -> sprintf "%s=-%s" (reg d) (reg r1) - | Osub, [r1;r2] -> sprintf "%s=%s-%s" (reg d) (reg r1) (reg r2) - | Omul, [r1;r2] -> sprintf "%s=%s*%s" (reg d) (reg r1) (reg r2) - | Omulimm n, [r1] -> sprintf "%s=%s*%ld" (reg d) (reg r1) (camlint_of_coqint n) - | Omulhs, [r1;r2] -> sprintf "%s=mulhs(%s,%s)" (reg d) (reg r1) (reg r2) - | Omulhu, [r1;r2] -> sprintf "%s=mulhu(%s,%s)" (reg d) (reg r1) (reg r2) - | Odiv, [r1;r2] -> sprintf "%s=%s /s %s" (reg d) (reg r1) (reg r2) - | Odivu, [r1;r2] -> sprintf "%s=%s /u %s" (reg d) (reg r1) (reg r2) - | Omod, [r1;r2] -> sprintf "%s=%s %%s %s" (reg d) (reg r1) (reg r2) - | Omodu, [r1;r2] -> sprintf "%s=%s %%u %s" (reg d) (reg r1) (reg r2) - | Oand, [r1;r2] -> sprintf "%s=%s & %s" (reg d) (reg r1) (reg r2) - | Oandimm n, [r1] -> sprintf "%s=%s & %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oor, [r1;r2] -> sprintf "%s=%s | %s" (reg d) (reg r1) (reg r2) - | Oorimm n, [r1] -> sprintf "%s=%s | %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oxor, [r1;r2] -> sprintf "%s=%s ^ %s" (reg d) (reg r1) (reg r2) - | Oxorimm n, [r1] -> sprintf "%s=%s ^ %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Onot, [r1] -> sprintf "%s=not(%s)" (reg d) (reg r1) - | Oshl, [r1;r2] -> sprintf "%s=%s << %s" (reg d) (reg r1) (reg r2) - | Oshlimm n, [r1] -> sprintf "%s=%s << %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshr, [r1;r2] -> sprintf "%s=%s >>s %s" (reg d) (reg r1) (reg r2) - | Oshrimm n, [r1] -> sprintf "%s=%s >>s %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrximm n, [r1] -> sprintf "%s=%s >>x %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshru, [r1;r2] -> sprintf "%s=%s >>u %s" (reg d) (reg r1) (reg r2) - | Oshruimm n, [r1] -> sprintf "%s=%s >>u %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Ororimm n, [r1] -> sprintf "%s=%s ror %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshldimm n, [r1;r2] -> sprintf "%s=(%s, %s) << %ld" (reg d) (reg r1) (reg r2) (camlint_of_coqint n) - | Olea addr, args -> sprintf "%s=addr" (reg d) - | Omakelong, [r1;r2] -> sprintf "%s=makelong(%s,%s)" (reg d) (reg r1) (reg r2) - | Olowlong, [r1] -> sprintf "%s=lowlong(%s)" (reg d) (reg r1) - | Ohighlong, [r1] -> sprintf "%s=highlong(%s)" (reg d) (reg r1) - | Ocast32signed, [r1] -> sprintf "%s=long32signed(%s)" (reg d) (reg r1) - | Ocast32unsigned, [r1] -> sprintf "%s=long32unsigned(%s)" (reg d) (reg r1) - | Onegl, [r1] -> sprintf "%s=-l %s" (reg d) (reg r1) - | Osubl, [r1;r2] -> sprintf "%s=%s -l %s" (reg d) (reg r1) (reg r2) - | Omull, [r1;r2] -> sprintf "%s=%s *l %s" (reg d) (reg r1) (reg r2) - | Omullimm n, [r1] -> sprintf "%s=%s *l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Omullhs, [r1;r2] -> sprintf "%s=mullhs(%s,%s)" (reg d) (reg r1) (reg r2) - | Omullhu, [r1;r2] -> sprintf "%s=mullhu(%s,%s)" (reg d) (reg r1) (reg r2) - | Odivl, [r1;r2] -> sprintf "%s=%s /ls %s" (reg d) (reg r1) (reg r2) - | Odivlu, [r1;r2] -> sprintf "%s=%s /lu %s" (reg d) (reg r1) (reg r2) - | Omodl, [r1;r2] -> sprintf "%s=%s %%ls %s" (reg d) (reg r1) (reg r2) - | Omodlu, [r1;r2] -> sprintf "%s=%s %%lu %s" (reg d) (reg r1) (reg r2) - | Oandl, [r1;r2] -> sprintf "%s=%s &l %s" (reg d) (reg r1) (reg r2) - | Oandlimm n, [r1] -> sprintf "%s=%s &l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Oorl, [r1;r2] -> sprintf "%s=%s |l %s" (reg d) (reg r1) (reg r2) - | Oorlimm n, [r1] -> sprintf "%s=%s |l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Oxorl, [r1;r2] -> sprintf "%s=%s ^l %s" (reg d) (reg r1) (reg r2) - | Oxorlimm n, [r1] -> sprintf "%s=%s ^l %Ld" (reg d) (reg r1) (camlint64_of_coqint n) - | Onotl, [r1] -> sprintf "%s=notl(%s)" (reg d) (reg r1) - | Oshll, [r1;r2] -> sprintf "%s=%s < sprintf "%s=%s < sprintf "%s=%s >>ls %s" (reg d) (reg r1) (reg r2) - | Oshrlimm n, [r1] -> sprintf "%s=%s >>ls %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrxlimm n, [r1] -> sprintf "%s=%s >>lx %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oshrlu, [r1;r2] -> sprintf "%s=%s >>lu %s" (reg d) (reg r1) (reg r2) - | Oshrluimm n, [r1] -> sprintf "%s=%s >>lu %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Ororlimm n, [r1] -> sprintf "%s=%s rorl %ld" (reg d) (reg r1) (camlint_of_coqint n) - | Oleal addr, args -> sprintf "%s=addr" (reg d) - | Onegf, [r1] -> sprintf "%s=negf(%s)" (reg d) (reg r1) - | Oabsf, [r1] -> sprintf "%s=absf(%s)" (reg d) (reg r1) - | Oaddf, [r1;r2] -> sprintf "%s=%s +f %s" (reg d) (reg r1) (reg r2) - | Osubf, [r1;r2] -> sprintf "%s=%s -f %s" (reg d) (reg r1) (reg r2) - | Omulf, [r1;r2] -> sprintf "%s=%s *f %s" (reg d) (reg r1) (reg r2) - | Odivf, [r1;r2] -> sprintf "%s=%s /f %s" (reg d) (reg r1) (reg r2) - | Onegfs, [r1] -> sprintf "%s=negfs(%s)" (reg d) (reg r1) - | Oabsfs, [r1] -> sprintf "%s=absfs(%s)" (reg d) (reg r1) - | Oaddfs, [r1;r2] -> sprintf "%s=%s +fs %s" (reg d) (reg r1) (reg r2) - | Osubfs, [r1;r2] -> sprintf "%s=%s -fs %s" (reg d) (reg r1) (reg r2) - | Omulfs, [r1;r2] -> sprintf "%s=%s *fs %s" (reg d) (reg r1) (reg r2) - | Odivfs, [r1;r2] -> sprintf "%s=%s /fs %s" (reg d) (reg r1) (reg r2) - | Osingleoffloat, [r1] -> sprintf "%s=singleoffloat(%s)" (reg d) (reg r1) - | Ofloatofsingle, [r1] -> sprintf "%s=floatofsingle(%s)" (reg d) (reg r1) - | Ointoffloat, [r1] -> sprintf "%s=intoffloat(%s)" (reg d) (reg r1) - | Ofloatofint, [r1] -> sprintf "%s=floatofint(%s)" (reg d) (reg r1) - | Ointofsingle, [r1] -> sprintf "%s=intofsingle(%s)" (reg d) (reg r1) - | Osingleofint, [r1] -> sprintf "%s=singleofint(%s)" (reg d) (reg r1) - | Olongoffloat, [r1] -> sprintf "%s=longoffloat(%s)" (reg d) (reg r1) - | Ofloatoflong, [r1] -> sprintf "%s=floatoflong(%s)" (reg d) (reg r1) - | Olongofsingle, [r1] -> sprintf "%s=longofsingle(%s)" (reg d) (reg r1) - | Osingleoflong, [r1] -> sprintf "%s=singleoflong(%s)" (reg d) (reg r1) - | Ocmp c, args -> sprintf "%s=cmp" (reg d) - | Osel (c, ty), r1::r2::args -> sprintf "%s=sel" (reg d) - | _, _ -> sprintf "N/a") - -module DFGDot = Graph.Graphviz.Dot(struct - let graph_attributes _ = [] - let default_vertex_attributes _ = [] - let vertex_name = function (i, instr) -> sprintf "\"%d:%s\"" i (print_instr instr) - let vertex_attributes _ = [] - let get_subgraph _ = None - let default_edge_attributes _ = [] - let edge_attributes _ = [] - - include DFG - end) - -module DFGDfs = Graph.Traverse.Dfs(DFG) - -module IMap = Map.Make (struct - type t = int - - let compare = compare -end) - -let gen_vertex instrs i = (i, List.nth instrs i) - -(** The DFG type defines a list of instructions with their data dependencies as [edges], which are - the pairs of integers that represent the index of the instruction in the [nodes]. The edges - always point from left to right. *) - -let print_list f out_chan a = - fprintf out_chan "[ "; - List.iter (fprintf out_chan "%a " f) a; - fprintf out_chan "]" - -let print_tuple out_chan a = - let l, r = a in - fprintf out_chan "(%d,%d)" l r - -(*let print_dfg out_chan dfg = - fprintf out_chan "{ nodes = %a, edges = %a }" - (print_list PrintRTLBlockInstr.print_bblock_body) - dfg.nodes (print_list print_tuple) dfg.edges*) - -let print_dfg = DFGDot.output_graph - -let read_process command = - let buffer_size = 2048 in - let buffer = Buffer.create buffer_size in - let string = Bytes.create buffer_size in - let in_channel = Unix.open_process_in command in - let chars_read = ref 1 in - while !chars_read <> 0 do - chars_read := input in_channel string 0 buffer_size; - Buffer.add_substring buffer (Bytes.to_string string) 0 !chars_read - done; - ignore (Unix.close_process_in in_channel); - Buffer.contents buffer - -let comb_delay = function - | RBnop -> 0 - | RBop (_, op, _, _) -> - (match op with - | Omove -> 0 - | Ointconst _ -> 0 - | Olongconst _ -> 0 - | Ocast8signed -> 0 - | Ocast8unsigned -> 0 - | Ocast16signed -> 0 - | Ocast16unsigned -> 0 - | Oneg -> 0 - | Onot -> 0 - | Oor -> 0 - | Oorimm _ -> 0 - | Oand -> 0 - | Oandimm _ -> 0 - | Oxor -> 0 - | Oxorimm _ -> 0 - | Omul -> 8 - | Omulimm _ -> 8 - | Omulhs -> 8 - | Omulhu -> 8 - | Odiv -> 72 - | Odivu -> 72 - | Omod -> 72 - | Omodu -> 72 - | _ -> 1) - | _ -> 1 - -let pipeline_stages = function - | RBload _ -> 2 - | RBop (_, op, _, _) -> - (match op with - | Odiv -> 32 - | Odivu -> 32 - | Omod -> 32 - | Omodu -> 32 - | _ -> 0) - | _ -> 0 - -(** Add a dependency if it uses a register that was written to previously. *) -let add_dep map i tree dfg curr = - match PTree.get curr tree with - | None -> dfg - | Some ip -> - let ipv = (List.nth map ip) in - 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. - - This function only gathers the RAW constraints, and will therefore only be active for operations - that modify registers, which is this case only affects loads and operations. *) -let accumulate_RAW_deps map dfg curr = - let i, dst_map, graph = dfg in - let acc_dep_instruction rs dst = - ( i + 1, - PTree.set dst i dst_map, - List.fold_left (add_dep map i dst_map) graph rs - ) - in - let acc_dep_instruction_nodst rs = - ( i + 1, - dst_map, - List.fold_left (add_dep map i dst_map) graph rs) - in - 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) - -(** Finds the next write to the [dst] register. This is a small optimisation so that only one - dependency is generated for a data dependency. *) -let rec find_next_dst_write i dst i' curr = - let check_dst dst' curr' = - if dst = dst' then Some (i, i') - else find_next_dst_write i dst (i' + 1) curr' - in - match curr with - | [] -> None - | RBop (_, _, _, dst') :: curr' -> check_dst dst' curr' - | RBload (_, _, _, _, dst') :: curr' -> check_dst dst' curr' - | _ :: curr' -> find_next_dst_write i dst (i' + 1) curr' - -let rec find_all_next_dst_read i dst i' curr = - let check_dst rs curr' = - if List.exists (fun x -> x = dst) rs - then (i, i') :: find_all_next_dst_read i dst (i' + 1) curr' - else find_all_next_dst_read i dst (i' + 1) curr' - in - match curr with - | [] -> [] - | RBop (_, _, rs, _) :: curr' -> check_dst rs curr' - | RBload (_, _, _, rs, _) :: curr' -> check_dst rs curr' - | RBstore (_, _, _, rs, src) :: curr' -> check_dst (src :: rs) curr' - | RBnop :: curr' -> find_all_next_dst_read i dst (i' + 1) curr' - | RBsetpred (_, _, rs, _) :: curr' -> check_dst rs curr' - -let drop i lst = - let rec drop' i' lst' = - match lst' with - | _ :: ls -> if i' = i then ls else drop' (i' + 1) ls - | [] -> [] - in - if i = 0 then lst else drop' 1 lst - -let take i lst = - let rec take' i' lst' = - match lst' with - | l :: ls -> if i' = i then [ l ] else l :: take' (i' + 1) ls - | [] -> [] - in - if i = 0 then [] else take' 1 lst - -let rec next_store i = function - | [] -> None - | RBstore (_, _, _, _, _) :: _ -> Some i - | _ :: rst -> next_store (i + 1) rst - -let rec next_load i = function - | [] -> None - | RBload (_, _, _, _, _) :: _ -> Some i - | _ :: rst -> next_load (i + 1) rst - -let accumulate_RAW_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBload (_, _, _, _, _) -> ( - 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) ) - | _ -> dfg - -let accumulate_WAR_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBstore (_, _, _, _, _) -> ( - match next_load 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) ) - | _ -> dfg - -let accumulate_WAW_mem_deps instrs dfg curri = - let i, curr = curri in - match curr with - | 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)) - | _ -> dfg - -(** Predicate dependencies. *) - -let rec in_predicate p p' = - match p' with - | Plit p'' -> P.to_int p = P.to_int (snd p'') - | Pand (p1, p2) -> in_predicate p p1 || in_predicate p p2 - | Por (p1, p2) -> in_predicate p p1 || in_predicate p p2 - | Ptrue -> false - | Pfalse -> false - -let get_predicate = function - | RBop (p, _, _, _) -> p - | RBload (p, _, _, _, _) -> p - | RBstore (p, _, _, _, _) -> p - | RBsetpred (p, _, _, _) -> p - | _ -> None - -let rec next_setpred p i = function - | [] -> None - | RBsetpred (_, _, _, p') :: rst -> - if in_predicate p' p then - Some i - else - next_setpred p (i + 1) rst - | _ :: rst -> next_setpred p (i + 1) rst - -let rec next_preduse p i instr= - let next p' rst = - if in_predicate p p' then - Some i - else - next_preduse p (i + 1) rst - in - match instr with - | [] -> None - | 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 = - let i, curr = curri in - match get_predicate curr with - | Some p -> ( - match next_setpred p 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) ) - | _ -> dfg - -let accumulate_WAR_pred_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBsetpred (_, _, _, p) -> ( - match next_preduse p 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) ) - | _ -> dfg - -let accumulate_WAW_pred_deps instrs dfg curri = - let i, curr = curri in - match curr with - | RBsetpred (_, _, _, p) -> ( - match next_setpred (Plit (true, p)) 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) ) - | _ -> dfg - -(** This function calculates the WAW dependencies, which happen when two writes are ordered one - after another and therefore have to be kept in that order. This accumulation might be redundant - if register renaming is done before hand, because then these dependencies can be avoided. *) -let accumulate_WAW_deps instrs dfg curri = - let i, curr = curri in - let dst_dep dst = - match find_next_dst_write i dst (i + 1) (drop (i + 1) instrs) with - | Some (a, b) -> DFG.add_edge dfg (gen_vertex instrs a) (gen_vertex instrs b) - | _ -> dfg - in - match curr with - | RBop (_, _, _, dst) -> dst_dep dst - | RBload (_, _, _, _, dst) -> dst_dep dst - | RBstore (_, _, _, _, _) -> ( - match next_store (i + 1) (drop (i + 1) instrs) with - | None -> dfg - | Some i' -> DFG.add_edge dfg (gen_vertex instrs i) (gen_vertex instrs i') ) - | _ -> dfg - -let accumulate_WAR_deps instrs dfg curri = - let i, curr = curri in - let dst_dep dst = - let dep_list = find_all_next_dst_read i dst 0 (take i instrs |> List.rev) - |> List.map (function (d, d') -> (i - d' - 1, d)) - in - List.fold_left (fun g -> - function (d, d') -> DFG.add_edge g (gen_vertex instrs d) (gen_vertex instrs d')) dfg dep_list - in - match curr with - | RBop (_, _, _, dst) -> dst_dep dst - | RBload (_, _, _, _, dst) -> dst_dep dst - | _ -> dfg - -let assigned_vars vars = function - | RBnop -> vars - | RBop (_, _, _, dst) -> dst :: vars - | RBload (_, _, _, _, dst) -> dst :: vars - | RBstore (_, _, _, _, _) -> vars - | RBsetpred (_, _, _, _) -> vars - -let get_pred = function - | RBnop -> None - | RBop (op, _, _, _) -> op - | RBload (op, _, _, _, _) -> op - | RBstore (op, _, _, _, _) -> op - | RBsetpred (_, _, _, _) -> None - -let independant_pred p p' = - match sat_pred_simple (Pand (p, p')) with - | None -> true - | _ -> false - -let check_dependent op1 op2 = - match op1, op2 with - | Some p, Some p' -> not (independant_pred p p') - | _, _ -> true - -let remove_unnecessary_deps graph = - let is_dependent v1 v2 g = - let (_, instr1) = v1 in - let (_, instr2) = v2 in - if check_dependent (get_pred instr1) (get_pred instr2) - then g - else DFG.remove_edge g v1 v2 - in - DFG.fold_edges is_dependent graph graph - -(** All the nodes in the DFG have to come after the source of the basic block, and should terminate - before the sink of the basic block. After that, there should be constraints for data - dependencies between nodes. *) -let gather_bb_constraints debug bb = - let ibody = List.mapi (fun i a -> (i, a)) bb.bb_body in - let dfg = List.fold_left (fun dfg v -> DFG.add_vertex dfg v) DFG.empty ibody in - let _, _, dfg' = - List.fold_left (accumulate_RAW_deps ibody) - (0, PTree.empty, dfg) - bb.bb_body - in - let dfg'' = List.fold_left (fun dfg f -> List.fold_left (f bb.bb_body) dfg ibody) dfg' - [ accumulate_WAW_deps; - accumulate_WAR_deps; - accumulate_RAW_mem_deps; - accumulate_WAR_mem_deps; - accumulate_WAW_mem_deps; - accumulate_RAW_pred_deps; - accumulate_WAR_pred_deps; - accumulate_WAW_pred_deps - ] - in - let dfg''' = remove_unnecessary_deps dfg'' in - (List.length bb.bb_body, dfg''', successors_instr bb.bb_exit) - -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 - then G.add_edge_e g (encode_bb n 0, 0, encode_var n (fst v1) 0) - else g) |> - (fun g' -> - if DFG.out_degree dfg v1 = 0 - 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, 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 = - List.fold_left (fun g n' -> - G.add_edge_e g (encode_bb n 1, -1, encode_bb n' 0) - ) constr succs - -module BFDFG = Graph.Path.BellmanFord(DFG)(struct - include DFG - type t = int - let weight = DFG.E.label - let compare = compare - let add = (+) - let zero = 0 - end) - -module TopoDFG = Graph.Topological.Make(DFG) - -let negate_graph constr = - DFG.fold_edges_e (function (v1, e, v2) -> fun g -> - DFG.add_edge_e g (v1, -e, v2) - ) constr DFG.empty - -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 - |> 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 maxf))) - 1), - encode_var n (fst v') 0) - ) constr (DFG.fold_vertex (fun v l -> - 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 = - | Mem - | SDiv - | UDiv - -type resources = { - res_mem: DFG.V.t list; - res_udiv: DFG.V.t list; - res_sdiv: DFG.V.t list; -} - -let find_resource = function - | RBload _ -> Some Mem - | RBstore _ -> Some Mem - | RBop (_, op, _, _) -> - ( match op with - | Odiv -> Some SDiv - | Odivu -> Some UDiv - | Omod -> Some SDiv - | Omodu -> Some UDiv - | _ -> None ) - | _ -> None - -let add_resource_constr n dfg constr = - let res = TopoDFG.fold (function (i, instr) -> - function {res_mem = ml; res_sdiv = sdl; res_udiv = udl} as r -> - match find_resource instr with - | Some SDiv -> {r with res_sdiv = (i, instr) :: sdl} - | Some UDiv -> {r with res_udiv = (i, instr) :: udl} - | Some Mem -> {r with res_mem = (i, instr) :: ml} - | None -> r - ) dfg {res_mem = []; res_sdiv = []; res_udiv = []} - in - let get_constraints l g = List.fold_left (fun gv v' -> - match gv with - | (g, None) -> (g, Some v') - | (g, Some v) -> - (G.add_edge_e g (encode_var n (fst v) 0, -1, encode_var n (fst v') 0), Some v') - ) (g, None) l |> fst - in - get_constraints (List.rev res.res_mem) constr - |> get_constraints (List.rev res.res_udiv) - |> get_constraints (List.rev res.res_sdiv) - -let gather_cfg_constraints c constr curr = - let (n, dfg) = curr in - match PTree.get (P.of_int n) c with - | 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 - |> List.filter (fun n' -> n' < n)) - |> add_cycle_constr 8 n dfg - |> add_resource_constr n dfg - -let rec intersperse s = function - | [] -> [] - | [ a ] -> [ a ] - | x :: xs -> x :: s :: intersperse s xs - -let print_objective constr = - let vars = G.fold_vertex (fun v1 l -> - match v1 with - | { sv_type = VarType _; sv_num = 0 } -> print_sv v1 :: l - | _ -> l - ) constr [] - in - "min: " ^ String.concat "" (intersperse " + " vars) ^ ";\n" - -let print_lp constr = - print_objective constr ^ - (G.fold_edges_e (function (e1, w, e2) -> fun s -> - s ^ sprintf "%s - %s <= %d;\n" (print_sv e1) (print_sv e2) w - ) constr "" |> - G.fold_vertex (fun v1 s -> - s ^ sprintf "int %s;\n" (print_sv v1) - ) constr) - -let update_schedule v = function Some l -> Some (v :: l) | None -> Some [ v ] - -let parse_soln (tree, bbtree) s = - let r = Str.regexp "var\\([0-9]+\\)n\\([0-9]+\\)_0[ ]+\\([0-9]+\\)" in - 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 (fn, oc) = Filename.open_temp_file "vericert_" "_lp_solve" in - fprintf oc "%s\n" (print_lp constr); - close_out oc; - - let res = Str.split (Str.regexp_string "\n") (read_process ("lp_solve " ^ fn)) - |> drop 3 - |> List.fold_left parse_soln (IMap.empty, 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 - List.fold_left (fun g v -> - List.fold_left (fun g' v' -> - let edges = DFG.find_all_edges dfg v v' in - List.fold_left DFG.add_edge_e g' edges - ) g l - ) dfg' l - -let rec all_successors dfg v = - List.concat (List.fold_left (fun l v -> - all_successors dfg v :: l - ) [] (DFG.succ dfg v)) - -let order_instr dfg = - DFG.fold_vertex (fun v li -> - if DFG.in_degree dfg v = 0 - then (List.map snd (v :: all_successors dfg v)) :: li - else li - ) dfg [] - -let combine_bb_schedule schedule s = - let i, st = s in - IMap.update st (update_schedule i) schedule - -(**let add_el dfg i l = - List.*) - -let check_in el = - List.exists (List.exists ((=) el)) - -let all_dfs dfg = - let roots = DFG.fold_vertex (fun v li -> - if DFG.in_degree dfg v = 0 then v :: li else li - ) dfg [] in - let dfg' = DFG.fold_edges (fun v1 v2 g -> DFG.add_edge g v2 v1) dfg dfg in - List.fold_left (fun a el -> - if check_in el a then a else - (DFGDfs.fold_component (fun v l -> v :: l) [] dfg' el) :: a) [] roots - -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 = - 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 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 - |> 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 -> - all_dfs x - |> List.map (subgraph x) - |> List.map (fun y -> - TopoDFG.fold (fun i l -> snd i :: l) y [] - |> List.rev))) body2 - (*|> (fun x -> TopoDFG.fold (fun i l -> snd i :: l) x []) - |> List.rev) body2*) - in - { bb_body = final_body2; - bb_exit = ctrl_flow - } - in - PTree.map f c - -let schedule entry (c : RTLBlock.bb RTLBlockInstr.code) = - let debug = true in - let transf_graph (_, dfg, _) = dfg in - let c' = PTree.map1 (fun x -> gather_bb_constraints false x |> transf_graph) c in - (*let _ = if debug then PTree.map (fun r o -> printf "##### %d #####\n%a\n\n" (P.to_int r) print_dfg o) c' else PTree.empty in*) - let cgraph = PTree.elements c' - |> List.map (function (x, y) -> (P.to_int x, y)) - |> List.fold_left (gather_cfg_constraints c) G.empty - in - let graph = open_out "constr_graph.dot" in - fprintf graph "%a\n" GDot.output_graph cgraph; - close_out graph; - let schedule' = solve_constraints cgraph in - (**IMap.iter (fun a b -> printf "##### %d #####\n%a\n\n" a (print_list print_tuple) b) schedule';*) - (**printf "Schedule: %a\n" (fun a x -> IMap.iter (fun d -> fprintf a "%d: %a\n" d (print_list print_tuple)) x) schedule';*) - transf_rtlpar c c' schedule' - -let rec find_reachable_states c e = - match PTree.get e c with - | Some { bb_exit = ex; _ } -> - e :: List.fold_left (fun x a -> List.concat [x; find_reachable_states c a]) [] - (successors_instr ex |> List.filter (fun x -> P.lt x e)) - | None -> assert false - -let add_to_tree c nt i = - match PTree.get i c with - | Some p -> PTree.set i p nt - | None -> assert false - -let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = - let scheduled = schedule f.fn_entrypoint f.fn_code in - let reachable = find_reachable_states scheduled f.fn_entrypoint - |> List.to_seq |> SS.of_seq |> SS.to_seq |> List.of_seq in - { fn_sig = f.fn_sig; - fn_params = f.fn_params; - fn_stacksize = f.fn_stacksize; - fn_code = scheduled (*List.fold_left (add_to_tree scheduled) PTree.empty reachable*); - fn_entrypoint = f.fn_entrypoint - } -#+end_src - -* RTLPargen -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargen.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpargen-main -#+begin_src coq -Require Import compcert.backend.Registers. -Require Import compcert.common.AST. -Require Import compcert.common.Globalenvs. -Require Import compcert.common.Memory. -Require Import compcert.common.Values. -Require Import compcert.lib.Floats. -Require Import compcert.lib.Integers. -Require Import compcert.lib.Maps. -Require compcert.verilog.Op. - -Require Import vericert.common.Vericertlib. -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLPar. -Require Import vericert.hls.RTLBlockInstr. -Require Import vericert.hls.Predicate. -Require Import vericert.hls.Abstr. -Import NE.NonEmptyNotation. - -#[local] Open Scope positive. -#[local] Open Scope forest. -#[local] Open Scope pred_op. -#+end_src - -** 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. - -#+name: rtlpargen-generation -#+begin_src coq -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) - end. - -Fixpoint replicate {A} (n: nat) (l: A) := - match n with - | O => nil - | S n => l :: replicate n l - end. - -Definition merge''' x y := - match x, y with - | Some p1, Some p2 => Some (Pand p1 p2) - | Some p, None | None, Some p => Some p - | None, None => None - end. - -Definition merge'' x := - match x with - | ((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 := - match pa, pf with - | (p, a), (p', f) => (merge''' p p', f a) - end. - -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. - -(*map (fun x => (fst x, Econs (snd x) Enil)) pel*) -Definition merge' (pel: pred_expr) (tpel: predicated expression_list) := - predicated_map (uncurry Econs) (predicated_prod pel tpel). - -Fixpoint merge (pel: list pred_expr): predicated expression_list := - match pel with - | nil => NE.singleton (T, Enil) - | a :: b => merge' a (merge b) - end. - -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 := - 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 := - 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 := - NE.map (fun x => (fst x, (snd x) pa pb pc)) pf. - -Definition predicated_from_opt {A: Type} (p: option pred_op) (a: A) := - match p with - | Some p' => NE.singleton (p', a) - | None => NE.singleton (T, a) - end. - -#[local] Open Scope non_empty_scope. -#[local] Open Scope pred_op. - -Fixpoint NEfold_left {A B} (f: A -> B -> A) (l: NE.non_empty B) (a: A) : A := - match l with - | NE.singleton a' => f a a' - | a' ::| b => NEfold_left f b (f a a') - end. - -Fixpoint NEapp {A} (l m: NE.non_empty A) := - match l with - | NE.singleton a => a ::| m - | a ::| b => a ::| NEapp b m - end. - -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. - -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). - -#+end_src - -*** Update Function -:PROPERTIES: -:CUSTOM_ID: update-function -:END: - -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. - -#+name: rtlpargen-update-function -#+begin_src coq -Definition update (f : forest) (i : instr) : forest := - match i with - | RBnop => f - | RBop p op rl r => - f # (Reg r) <- - (app_predicated p - (f # (Reg r)) - (map_predicated (pred_ret (Eop op)) (merge (list_translation rl f)))) - | RBload p chunk addr rl r => - f # (Reg r) <- - (app_predicated p - (f # (Reg r)) - (map_predicated - (map_predicated (pred_ret (Eload chunk addr)) (merge (list_translation rl f))) - (f # Mem))) - | RBstore p chunk addr rl r => - f # Mem <- - (app_predicated p - (f # Mem) - (map_predicated - (map_predicated - (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)))) - end. -#+end_src - -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. - -#+name: rtlpargen-abstract-seq -#+begin_src coq -Fixpoint abstract_sequence (f : forest) (b : list instr) : forest := - match b with - | nil => f - | i :: l => abstract_sequence (update f i) l - end. -#+end_src - -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. - -#+name: rtlpargen-check-functions -#+begin_src coq -Definition check_control_flow_instr (c1 c2: cf_instr) : bool := - if cf_instr_eq c1 c2 then true else false. -#+end_src - -We define the top-level oracle that will check if two basic blocks are equivalent after a -scheduling transformation. - -#+name: rtlpargen-top-check-functions -#+begin_src coq -Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := - match bb with - | nil => - match bbt with - | nil => true - | _ => false - end - | _ => true - end. - -Definition schedule_oracle (bb: RTLBlock.bblock) (bbt: RTLPar.bblock) : bool := - check (abstract_sequence empty (bb_body bb)) - (abstract_sequence empty (concat (concat (bb_body bbt)))) && - check_control_flow_instr (bb_exit bb) (bb_exit bbt) && - empty_trees (bb_body bb) (bb_body bbt). - -Definition check_scheduled_trees := beq2 schedule_oracle. - -Ltac solve_scheduled_trees_correct := - intros; unfold check_scheduled_trees in *; - match goal with - | [ H: context[beq2 _ _ _], x: positive |- _ ] => - rewrite beq2_correct in H; specialize (H x) - end; repeat destruct_match; crush. - -Lemma check_scheduled_trees_correct: - forall f1 f2 x y1, - check_scheduled_trees f1 f2 = true -> - PTree.get x f1 = Some y1 -> - exists y2, PTree.get x f2 = Some y2 /\ schedule_oracle y1 y2 = true. -Proof. solve_scheduled_trees_correct; eexists; crush. Qed. - -Lemma check_scheduled_trees_correct2: - forall f1 f2 x, - check_scheduled_trees f1 f2 = true -> - PTree.get x f1 = None -> - PTree.get x f2 = None. -Proof. solve_scheduled_trees_correct. Qed. - -#+end_src - -** Top-level Functions - -#+name: rtlpargen-top-level-functions -#+begin_src coq -Parameter schedule : RTLBlock.function -> 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) - f.(fn_params) - f.(fn_stacksize) - tfcode - f.(fn_entrypoint)) - else - 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. -#+end_src - -* RTLPargenproof -:PROPERTIES: -:header-args:coq: :comments noweb :noweb no-export :padline yes :tangle ../src/hls/RTLPargenproof.v -:END: - -#+begin_src coq :comments no :padline no :exports none -<> -#+end_src - -#+name: rtlpargenproof-imports -#+begin_src coq -Require Import compcert.backend.Registers. -Require Import compcert.common.AST. -Require Import compcert.common.Errors. -Require Import compcert.common.Linking. -Require Import compcert.common.Globalenvs. -Require Import compcert.common.Memory. -Require Import compcert.common.Values. -Require Import compcert.lib.Maps. - -Require Import vericert.common.Vericertlib. -Require Import vericert.hls.RTLBlock. -Require Import vericert.hls.RTLPar. -Require Import vericert.hls.RTLBlockInstr. -Require Import vericert.hls.RTLPargen. -Require Import vericert.hls.Predicate. -Require Import vericert.hls.Abstr. - -#[local] Open Scope positive. -#[local] Open Scope forest. -#[local] Open Scope pred_op. -#+end_src - -** RTLBlock to abstract translation - -Correctness of translation from RTLBlock to the abstract interpretation language. - -#+name: rtlpargenproof-main -#+begin_src coq -(*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. - -Inductive state_lessdef : instr_state -> instr_state -> Prop := - state_lessdef_intro : - forall rs1 rs2 m1, - (forall x, rs1 !! x = rs2 !! x) -> - state_lessdef (mk_instr_state rs1 m1) (mk_instr_state rs2 m1). - -Ltac inv_simp := - repeat match goal with - | H: exists _, _ |- _ => inv H - end; simplify. - -*) - -Definition check_dest i r' := - match i with - | RBop p op rl r => (r =? r')%positive - | RBload p chunk addr rl r => (r =? r')%positive - | _ => false - end. - -Lemma check_dest_dec i r : {check_dest i r = true} + {check_dest i r = false}. -Proof. destruct (check_dest i r); tauto. Qed. - -Fixpoint check_dest_l l r := - match l with - | nil => false - | a :: b => check_dest a r || check_dest_l b r - end. - -Lemma check_dest_l_forall : - forall l r, - check_dest_l l r = false -> - Forall (fun x => check_dest x r = false) l. -Proof. induction l; crush. Qed. - -Lemma check_dest_l_dec i r : {check_dest_l i r = true} + {check_dest_l i r = false}. -Proof. destruct (check_dest_l i r); tauto. Qed. - -Lemma check_dest_update : - forall f i r, - check_dest i r = false -> - (update f i) # (Reg r) = f # (Reg r). -Proof. - destruct i; crush; try apply Pos.eqb_neq in H; apply genmap1; crush. -Qed. - -Lemma check_dest_l_forall2 : - forall l r, - Forall (fun x => check_dest x r = false) l -> - check_dest_l l r = false. -Proof. - induction l; crush. - inv H. apply orb_false_intro; crush. -Qed. - -Lemma check_dest_l_ex2 : - forall l r, - (exists a, In a l /\ check_dest a r = true) -> - check_dest_l l r = true. -Proof. - induction l; crush. - specialize (IHl r). inv H. - apply orb_true_intro; crush. - apply orb_true_intro; crush. - right. apply IHl. exists x. auto. -Qed. - -Lemma check_list_l_false : - forall l x r, - check_dest_l (l ++ x :: nil) r = false -> - check_dest_l l r = false /\ check_dest x r = false. -Proof. - simplify. - apply check_dest_l_forall in H. apply Forall_app in H. - simplify. apply check_dest_l_forall2; auto. - apply check_dest_l_forall in H. apply Forall_app in H. - simplify. inv H1. auto. -Qed. - -Lemma check_dest_l_ex : - forall l r, - check_dest_l l r = true -> - exists a, In a l /\ check_dest a r = true. -Proof. - induction l; crush. - destruct (check_dest a r) eqn:?; try solve [econstructor; crush]. - simplify. - exploit IHl. apply H. simplify. econstructor. simplify. right. eassumption. - auto. -Qed. - -Lemma check_list_l_true : - forall l x r, - check_dest_l (l ++ x :: nil) r = true -> - check_dest_l l r = true \/ check_dest x r = true. -Proof. - simplify. - apply check_dest_l_ex in H; simplify. - apply in_app_or in H. inv H. left. - apply check_dest_l_ex2. exists x0. auto. - inv H0; auto. -Qed. - -Lemma check_dest_l_dec2 l r : - {Forall (fun x => check_dest x r = false) l} - + {exists a, In a l /\ check_dest a r = true}. -Proof. - destruct (check_dest_l_dec l r); [right | left]; - auto using check_dest_l_ex, check_dest_l_forall. -Qed. - -Lemma abstr_comp : - forall l i f x x0, - abstract_sequence f (l ++ i :: nil) = x -> - abstract_sequence f l = x0 -> - x = update x0 i. -Proof. induction l; intros; crush; eapply IHl; eauto. Qed. - -(* - -Lemma gen_list_base: - forall FF ge sp l rs exps st1, - (forall x, @sem_value FF ge sp st1 (exps # (Reg x)) (rs !! x)) -> - sem_val_list ge sp st1 (list_translation l exps) rs ## l. -Proof. - induction l. - intros. simpl. constructor. - intros. simpl. eapply Scons; eauto. -Qed. - -Lemma check_dest_update2 : - forall f r rl op p, - (update f (RBop p op rl r)) # (Reg r) = Eop op (list_translation rl f) (f # Mem). -Proof. crush; rewrite map2; auto. Qed. - -Lemma check_dest_update3 : - forall f r rl p addr chunk, - (update f (RBload p chunk addr rl r)) # (Reg r) = Eload chunk addr (list_translation rl f) (f # Mem). -Proof. crush; rewrite map2; auto. Qed. - -Lemma abstract_seq_correct_aux: - forall FF ge sp i st1 st2 st3 f, - @step_instr FF ge sp st3 i st2 -> - sem ge sp st1 f st3 -> - sem ge sp st1 (update f i) st2. -Proof. - intros; inv H; simplify. - { simplify; eauto. } (*apply match_states_refl. }*) - { inv H0. inv H6. destruct st1. econstructor. simplify. - constructor. intros. - destruct (resource_eq (Reg res) (Reg x)). inv e. - rewrite map2. econstructor. eassumption. apply gen_list_base; eauto. - rewrite Regmap.gss. eauto. - assert (res <> x). { unfold not in *. intros. apply n. rewrite H0. auto. } - rewrite Regmap.gso by auto. - rewrite genmap1 by auto. auto. - - rewrite genmap1; crush. } - { inv H0. inv H7. constructor. constructor. intros. - destruct (Pos.eq_dec dst x); subst. - rewrite map2. econstructor; eauto. - apply gen_list_base. auto. rewrite Regmap.gss. auto. - rewrite genmap1. rewrite Regmap.gso by auto. auto. - unfold not in *; intros. inv H0. auto. - rewrite genmap1; crush. - } - { inv H0. inv H7. constructor. constructor; intros. - rewrite genmap1; crush. - rewrite map2. econstructor; eauto. - apply gen_list_base; auto. - } -Qed. - -Lemma regmap_list_equiv : - forall A (rs1: Regmap.t A) rs2, - (forall x, rs1 !! x = rs2 !! x) -> - forall rl, rs1##rl = rs2##rl. -Proof. induction rl; crush. Qed. - -Lemma sem_update_Op : - forall A ge sp st f st' r l o0 o m rs v, - @sem A ge sp st f st' -> - Op.eval_operation ge sp o0 rs ## l m = Some v -> - match_states st' (mk_instr_state rs m) -> - exists tst, - sem ge sp st (update f (RBop o o0 l r)) tst /\ match_states (mk_instr_state (Regmap.set r v rs) m) tst. -Proof. - intros. inv H1. simplify. - destruct st. - econstructor. simplify. - { constructor. - { constructor. intros. destruct (Pos.eq_dec x r); subst. - { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. - { inv H9. eapply gen_list_base; eauto. } - { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } - { inv H. rewrite genmap1; crush. eauto. } } - { constructor; eauto. intros. - destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed. - -Lemma sem_update_load : - forall A ge sp st f st' r o m a l m0 rs v a0, - @sem A ge sp st f st' -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.loadv m m0 a0 = Some v -> - match_states st' (mk_instr_state rs m0) -> - exists tst : instr_state, - sem ge sp st (update f (RBload o m a l r)) tst - /\ match_states (mk_instr_state (Regmap.set r v rs) m0) tst. -Proof. - intros. inv H2. pose proof H. inv H. inv H9. - destruct st. - econstructor; simplify. - { constructor. - { constructor. intros. - destruct (Pos.eq_dec x r); subst. - { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. - rewrite <- H6. eauto. - instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. auto. } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. } } - { rewrite genmap1; crush. eauto. } } - { constructor; auto; intros. destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed. - -Lemma sem_update_store : - forall A ge sp a0 m a l r o f st m' rs m0 st', - @sem A ge sp st f st' -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.storev m m0 a0 rs !! r = Some m' -> - match_states st' (mk_instr_state rs m0) -> - exists tst, sem ge sp st (update f (RBstore o m a l r)) tst - /\ match_states (mk_instr_state rs m') tst. -Proof. - intros. inv H2. pose proof H. inv H. inv H9. - destruct st. - econstructor; simplify. - { econstructor. - { econstructor; intros. rewrite genmap1; crush. } - { rewrite map2. econstructor; eauto. eapply gen_list_base. intros. rewrite <- H6. - eauto. specialize (H6 r). rewrite H6. eauto. } } - { econstructor; eauto. } -Qed. - -Lemma sem_update : - forall A ge sp st x st' st'' st''' f, - sem ge sp st f st' -> - match_states st' st''' -> - @step_instr A ge sp st''' x st'' -> - exists tst, sem ge sp st (update f x) tst /\ match_states st'' tst. -Proof. - intros. destruct x; inv H1. - { econstructor. split. - apply sem_update_RBnop. eassumption. - apply match_states_commut. auto. } - { eapply sem_update_Op; eauto. } - { eapply sem_update_load; eauto. } - { eapply sem_update_store; eauto. } -Qed. - -Lemma sem_update2_Op : - forall A ge sp st f r l o0 o m rs v, - @sem A ge sp st f (mk_instr_state rs m) -> - Op.eval_operation ge sp o0 rs ## l m = Some v -> - sem ge sp st (update f (RBop o o0 l r)) (mk_instr_state (Regmap.set r v rs) m). -Proof. - intros. destruct st. constructor. - inv H. inv H6. - { constructor; intros. simplify. - destruct (Pos.eq_dec r x); subst. - { rewrite map2. econstructor. eauto. - apply gen_list_base. eauto. - rewrite Regmap.gss. auto. } - { rewrite genmap1; crush. rewrite Regmap.gso; auto. } } - { simplify. rewrite genmap1; crush. inv H. eauto. } -Qed. - -Lemma sem_update2_load : - forall A ge sp st f r o m a l m0 rs v a0, - @sem A ge sp st f (mk_instr_state rs m0) -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.loadv m m0 a0 = Some v -> - sem ge sp st (update f (RBload o m a l r)) (mk_instr_state (Regmap.set r v rs) m0). -Proof. - intros. simplify. inv H. inv H7. constructor. - { constructor; intros. destruct (Pos.eq_dec r x); subst. - { rewrite map2. rewrite Regmap.gss. econstructor; eauto. - apply gen_list_base; eauto. } - { rewrite genmap1; crush. rewrite Regmap.gso; eauto. } - } - { simplify. rewrite genmap1; crush. } -Qed. - -Lemma sem_update2_store : - forall A ge sp a0 m a l r o f st m' rs m0, - @sem A ge sp st f (mk_instr_state rs m0) -> - Op.eval_addressing ge sp a rs ## l = Some a0 -> - Mem.storev m m0 a0 rs !! r = Some m' -> - sem ge sp st (update f (RBstore o m a l r)) (mk_instr_state rs m'). -Proof. - intros. simplify. inv H. inv H7. constructor; simplify. - { econstructor; intros. rewrite genmap1; crush. } - { rewrite map2. econstructor; eauto. apply gen_list_base; eauto. } -Qed. - -Lemma sem_update2 : - forall A ge sp st x st' st'' f, - sem ge sp st f st' -> - @step_instr A ge sp st' x st'' -> - sem ge sp st (update f x) st''. -Proof. - intros. - destruct x; inv H0; - eauto using sem_update_RBnop, sem_update2_Op, sem_update2_load, sem_update2_store. -Qed. - -Lemma abstr_sem_val_mem : - forall A B ge tge st tst sp a, - ge_preserved ge tge -> - forall v m, - (@sem_mem A ge sp st a m /\ match_states st tst -> @sem_mem B tge sp tst a m) /\ - (@sem_value A ge sp st a v /\ match_states st tst -> @sem_value B tge sp tst a v). -Proof. - intros * H. - apply expression_ind2 with - - (P := fun (e1: expression) => - forall v m, - (@sem_mem A ge sp st e1 m /\ match_states st tst -> @sem_mem B tge sp tst e1 m) /\ - (@sem_value A ge sp st e1 v /\ match_states st tst -> @sem_value B tge sp tst e1 v)) - - (P0 := fun (e1: expression_list) => - forall lv, @sem_val_list A ge sp st e1 lv /\ match_states st tst -> @sem_val_list B tge sp tst e1 lv); - simplify; intros; simplify. - { inv H1. inv H2. constructor. } - { inv H2. inv H1. rewrite H0. constructor. } - { inv H3. } - { inv H3. inv H4. econstructor. apply H1; auto. simplify. eauto. constructor. auto. auto. - apply H0; simplify; eauto. constructor; eauto. - unfold ge_preserved in *. simplify. rewrite <- H2. auto. - } - { inv H3. } - { inv H3. inv H4. econstructor. apply H1; eauto; simplify; eauto. constructor; eauto. - apply H0; simplify; eauto. constructor; eauto. - inv H. rewrite <- H4. eauto. - auto. - } - { inv H4. inv H5. econstructor. apply H0; eauto. simplify; eauto. constructor; eauto. - apply H2; eauto. simplify; eauto. constructor; eauto. - apply H1; eauto. simplify; eauto. constructor; eauto. - inv H. rewrite <- H5. eauto. auto. - } - { inv H4. } - { inv H1. constructor. } - { inv H3. constructor; auto. apply H0; eauto. apply Mem.empty. } -Qed. - -Lemma abstr_sem_value : - forall a A B ge tge sp st tst v, - @sem_value A ge sp st a v -> - ge_preserved ge tge -> - match_states st tst -> - @sem_value B tge sp tst a v. -Proof. intros; eapply abstr_sem_val_mem; eauto; apply Mem.empty. Qed. - -Lemma abstr_sem_mem : - forall a A B ge tge sp st tst v, - @sem_mem A ge sp st a v -> - ge_preserved ge tge -> - match_states st tst -> - @sem_mem B tge sp tst a v. -Proof. intros; eapply abstr_sem_val_mem; eauto. Qed. - -Lemma abstr_sem_regset : - forall a a' A B ge tge sp st tst rs, - @sem_regset A ge sp st a rs -> - ge_preserved ge tge -> - (forall x, a # x = a' # x) -> - match_states st tst -> - exists rs', @sem_regset B tge sp tst a' rs' /\ (forall x, rs !! x = rs' !! x). -Proof. - inversion 1; intros. - inv H7. - econstructor. simplify. econstructor. intros. - eapply abstr_sem_value; eauto. rewrite <- H6. - eapply H0. constructor; eauto. - auto. -Qed. - -Lemma abstr_sem : - forall a a' A B ge tge sp st tst st', - @sem A ge sp st a st' -> - ge_preserved ge tge -> - (forall x, a # x = a' # x) -> - match_states st tst -> - exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. -Proof. - inversion 1; subst; intros. - inversion H4; subst. - exploit abstr_sem_regset; eauto; inv_simp. - do 3 econstructor; eauto. - rewrite <- H3. - eapply abstr_sem_mem; eauto. -Qed. - -Lemma abstract_execution_correct': - forall A B ge tge sp st' a a' st tst, - @sem A ge sp st a st' -> - ge_preserved ge tge -> - check a a' = true -> - match_states st tst -> - exists tst', @sem B tge sp tst a' tst' /\ match_states st' tst'. -Proof. - intros; - pose proof (check_correct a a' H1); - eapply abstr_sem; eauto. -Qed. - -Lemma states_match : - forall st1 st2 st3 st4, - match_states st1 st2 -> - match_states st2 st3 -> - match_states st3 st4 -> - match_states st1 st4. -Proof. - intros * H1 H2 H3; destruct st1; destruct st2; destruct st3; destruct st4. - inv H1. inv H2. inv H3; constructor. - unfold regs_lessdef in *. intros. - repeat match goal with - | H: forall _, _, r : positive |- _ => specialize (H r) - end. - congruence. - auto. -Qed. - -Lemma step_instr_block_same : - forall ge sp st st', - step_instr_block ge sp st nil st' -> - st = st'. -Proof. inversion 1; auto. Qed. - -Lemma step_instr_seq_same : - forall ge sp st st', - step_instr_seq ge sp st nil st' -> - st = st'. -Proof. inversion 1; auto. Qed. - -Lemma sem_update' : - forall A ge sp st a x st', - sem ge sp st (update (abstract_sequence empty a) x) st' -> - exists st'', - @step_instr A ge sp st'' x st' /\ - sem ge sp st (abstract_sequence empty a) st''. -Proof. - Admitted. - -Lemma rtlpar_trans_correct : - forall bb ge sp sem_st' sem_st st, - sem ge sp sem_st (abstract_sequence empty (concat (concat bb))) sem_st' -> - match_states sem_st st -> - exists st', RTLPar.step_instr_block ge sp st bb st' - /\ match_states sem_st' st'. -Proof. - induction bb using rev_ind. - { repeat econstructor. eapply abstract_interp_empty3 in H. - inv H. inv H0. constructor; congruence. } - { simplify. inv H0. repeat rewrite concat_app in H. simplify. - rewrite app_nil_r in H. - exploit sem_separate; eauto; inv_simp. - repeat econstructor. admit. admit. - } -Admitted. - -(*Lemma abstract_execution_correct_ld: - forall bb bb' cfi ge tge sp st st' tst, - RTLBlock.step_instr_list ge sp st bb st' -> - ge_preserved ge tge -> - schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi) = true -> - match_states_ld st tst -> - exists tst', RTLPar.step_instr_block tge sp tst bb' tst' - /\ match_states st' tst'. -Proof. - intros.*) -*) - -Lemma match_states_list : - forall A (rs: Regmap.t A) rs', - (forall r, rs !! r = rs' !! r) -> - forall l, rs ## l = rs' ## l. -Proof. induction l; crush. Qed. - -Lemma PTree_matches : - forall A (v: A) res rs rs', - (forall r, rs !! r = rs' !! r) -> - forall x, (Regmap.set res v rs) !! x = (Regmap.set res v rs') !! x. -Proof. - intros; destruct (Pos.eq_dec x res); subst; - [ repeat rewrite Regmap.gss by auto - | repeat rewrite Regmap.gso by auto ]; auto. -Qed. - -Lemma abstract_interp_empty3 : - forall A ctx st', - @sem A ctx empty st' -> match_states (ctx_is ctx) st'. -Proof. - inversion 1; subst; simplify. destruct ctx. - destruct ctx_is. - constructor; intros. - - inv H0. specialize (H3 x). inv H3. inv H8. reflexivity. - - inv H1. specialize (H3 x). inv H3. inv H8. reflexivity. - - inv H2. inv H8. reflexivity. -Qed. - -Lemma step_instr_matches : - forall A a ge sp st st', - @step_instr A ge sp st a st' -> - forall tst, - match_states st tst -> - exists tst', step_instr ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction 1; simplify; - match goal with H: match_states _ _ |- _ => inv H end; - try solve [repeat econstructor; try erewrite match_states_list; - try apply PTree_matches; eauto; - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end]. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - apply PTree_matches; assumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - econstructor; auto. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - apply PTree_matches; assumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - erewrite <- eval_predf_pr_equiv; eassumption. - econstructor; auto. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - - destruct p. match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - erewrite <- eval_predf_pr_equiv; eassumption. - repeat (econstructor; try apply eval_pred_false); eauto. try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - erewrite <- eval_predf_pr_equiv; eassumption. - match goal with H: eval_pred _ _ _ _ |- _ => inv H end. - repeat econstructor; try erewrite match_states_list; eauto. - match goal with - H: forall _, _ |- context[Mem.storev] => erewrite <- H; eauto - end. - - admit. Admitted. - -Lemma step_instr_list_matches : - forall a ge sp st st', - step_instr_list ge sp st a st' -> - forall tst, match_states st tst -> - exists tst', step_instr_list ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction a; intros; inv H; - try (exploit step_instr_matches; eauto; []; simplify; - exploit IHa; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma step_instr_seq_matches : - forall a ge sp st st', - step_instr_seq ge sp st a st' -> - forall tst, match_states st tst -> - exists tst', step_instr_seq ge sp tst a tst' - /\ match_states st' tst'. -Proof. - induction a; intros; inv H; - try (exploit step_instr_list_matches; eauto; []; simplify; - exploit IHa; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma step_instr_block_matches : - forall bb ge sp st st', - step_instr_block ge sp st bb st' -> - forall tst, match_states st tst -> - exists tst', step_instr_block ge sp tst bb tst' - /\ match_states st' tst'. -Proof. - induction bb; intros; inv H; - try (exploit step_instr_seq_matches; eauto; []; simplify; - exploit IHbb; eauto; []; simplify); repeat econstructor; eauto. -Qed. - -Lemma rtlblock_trans_correct' : - forall bb ge sp st x st'', - RTLBlock.step_instr_list ge sp st (bb ++ x :: nil) st'' -> - exists st', RTLBlock.step_instr_list ge sp st bb st' - /\ step_instr ge sp st' x st''. -Proof. - induction bb. - crush. exists st. - split. constructor. inv H. inv H6. auto. - crush. inv H. exploit IHbb. eassumption. simplify. - econstructor. split. - econstructor; eauto. eauto. -Qed. - -Lemma abstract_interp_empty A st : @sem A st empty (ctx_is st). -Proof. destruct st, ctx_is. simpl. repeat econstructor. Qed. - -Lemma abstract_seq : - forall l f i, - abstract_sequence f (l ++ i :: nil) = update (abstract_sequence f l) i. -Proof. induction l; crush. Qed. - -Lemma abstract_sequence_update : - forall l r f, - check_dest_l l r = false -> - (abstract_sequence f l) # (Reg r) = f # (Reg r). -Proof. - induction l using rev_ind; crush. - rewrite abstract_seq. rewrite check_dest_update. apply IHl. - apply check_list_l_false in H. tauto. - apply check_list_l_false in H. tauto. -Qed. - -(*Lemma sem_separate : - forall A ctx b a st', - sem ctx (abstract_sequence empty (a ++ b)) st' -> - exists st'', - @sem A ctx (abstract_sequence empty a) st'' - /\ @sem A (mk_ctx st'' (ctx_sp ctx) (ctx_ge ctx)) (abstract_sequence empty b) st'. -Proof. - induction b using rev_ind; simplify. - { econstructor. simplify. rewrite app_nil_r in H. eauto. apply abstract_interp_empty. } - { simplify. rewrite app_assoc in H. rewrite abstract_seq in H. - exploit sem_update'; eauto; simplify. - exploit IHb; eauto; inv_simp. - econstructor; split; eauto. - rewrite abstract_seq. - eapply sem_update2; eauto. - } -Qed.*) - -Lemma sem_update_RBnop : - forall A ctx f st', - @sem A ctx f st' -> sem ctx (update f RBnop) st'. -Proof. auto. Qed. - -Lemma sem_update_Op : - forall A ge sp ist f st' r l o0 o m rs v ps, - @sem A (mk_ctx ist sp ge) f st' -> - eval_predf ps o = true -> - Op.eval_operation ge sp o0 (rs ## l) m = Some v -> - match_states st' (mk_instr_state rs ps m) -> - exists tst, - sem (mk_ctx ist sp ge) (update f (RBop (Some o) o0 l r)) tst - /\ match_states (mk_instr_state (Regmap.set r v rs) ps m) tst. -Proof. - intros. inv H1. inv H. inv H1. inv H3. simplify. - econstructor. simplify. - { constructor; try constructor; intros; try solve [rewrite genmap1; now eauto]. - destruct (Pos.eq_dec x r); subst. - { rewrite map2. specialize (H1 r). inv H1. -(*} - } - destruct st. - econstructor. simplify. - { constructor. - { constructor. intros. destruct (Pos.eq_dec x r); subst. - { pose proof (H5 r). rewrite map2. pose proof H. inv H. econstructor; eauto. - { inv H9. eapply gen_list_base; eauto. } - { instantiate (1 := (Regmap.set r v rs0)). rewrite Regmap.gss. erewrite regmap_list_equiv; eauto. } } - { rewrite Regmap.gso by auto. rewrite genmap1; crush. inv H. inv H7; eauto. } } - { inv H. rewrite genmap1; crush. eauto. } } - { constructor; eauto. intros. - destruct (Pos.eq_dec r x); - subst; [repeat rewrite Regmap.gss | repeat rewrite Regmap.gso]; auto. } -Qed.*) Admitted. - -Lemma sem_update : - forall A ge sp st x st' st'' st''' f, - sem (mk_ctx st sp ge) f st' -> - 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. - 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', - RTLBlock.step_instr_list ge sp st bb st' -> - forall tst, - match_states st tst -> - exists tst', sem (mk_ctx tst sp ge) (abstract_sequence empty bb) tst' - /\ match_states st' tst'. -Proof. - induction bb using rev_ind; simplify. - { econstructor. simplify. apply abstract_interp_empty. - inv H. auto. } - { apply rtlblock_trans_correct' in H. simplify. - rewrite abstract_seq. - exploit IHbb; try eassumption; []; simplify. - exploit sem_update. apply H1. symmetry; eassumption. - eauto. simplify. econstructor. split. apply H3. - auto. } -Qed. - -Lemma abstract_execution_correct: - forall bb bb' cfi cfi' ge tge sp st st' tst, - RTLBlock.step_instr_list ge sp st bb st' -> - ge_preserved ge tge -> - schedule_oracle (mk_bblock bb cfi) (mk_bblock bb' cfi') = true -> - match_states st tst -> - exists tst', RTLPar.step_instr_block tge sp tst bb' tst' - /\ match_states st' tst'. -Proof. - intros. - unfold schedule_oracle in *. simplify. unfold empty_trees in H4. - exploit rtlblock_trans_correct; try eassumption; []; simplify. -(*) exploit abstract_execution_correct'; - try solve [eassumption | apply state_lessdef_match_sem; eassumption]. - apply match_states_commut. eauto. inv_simp. - exploit rtlpar_trans_correct; try eassumption; []; inv_simp. - exploit step_instr_block_matches; eauto. apply match_states_commut; eauto. inv_simp. - repeat match goal with | H: match_states _ _ |- _ => inv H end. - do 2 econstructor; eauto. - econstructor; congruence. -Qed.*)Admitted. - -Definition match_prog (prog : RTLBlock.program) (tprog : RTLPar.program) := - match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq prog tprog. - -Inductive match_stackframes: RTLBlock.stackframe -> RTLPar.stackframe -> Prop := -| match_stackframe: - forall f tf res sp pc rs rs' ps ps', - transl_function f = OK tf -> - (forall x, rs !! x = rs' !! x) -> - (forall x, ps !! x = ps' !! x) -> - match_stackframes (Stackframe res f sp pc rs ps) - (Stackframe res tf sp pc rs' ps'). - -Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop := -| match_state: - forall sf f sp pc rs rs' m sf' tf ps ps' - (TRANSL: transl_function f = OK tf) - (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_returnstate: - forall stack stack' v m - (STACKS: list_forall2 match_stackframes stack stack'), - match_states (Returnstate stack v m) - (Returnstate stack' v m) -| match_callstate: - forall stack stack' f tf args m - (TRANSL: transl_fundef f = OK tf) - (STACKS: list_forall2 match_stackframes stack stack'), - match_states (Callstate stack f args m) - (Callstate stack' tf args m). - -Section CORRECTNESS. - - Context (prog: RTLBlock.program) (tprog : RTLPar.program). - Context (TRANSL: match_prog prog tprog). - - Let ge : RTLBlock.genv := Globalenvs.Genv.globalenv prog. - Let tge : RTLPar.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. - Hint Resolve symbols_preserved : rtlgp. - - Lemma function_ptr_translated: - forall (b: Values.block) (f: RTLBlock.fundef), - Genv.find_funct_ptr ge b = Some f -> - exists tf, - Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = Errors.OK tf. - Proof using TRANSL. - intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto. - intros (cu & tf & P & Q & R); exists tf; auto. - Qed. - - Lemma functions_translated: - forall (v: Values.val) (f: RTLBlock.fundef), - Genv.find_funct ge v = Some f -> - exists tf, - Genv.find_funct tge v = Some tf /\ transl_fundef f = Errors.OK tf. - Proof using TRANSL. - intros. exploit (Genv.find_funct_match TRANSL); eauto. - intros (cu & tf & P & Q & R); exists tf; auto. - Qed. - - Lemma senv_preserved: - Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). - Proof (Genv.senv_transf_partial TRANSL). - Hint Resolve senv_preserved : rtlgp. - - Lemma sig_transl_function: - forall (f: RTLBlock.fundef) (tf: RTLPar.fundef), - transl_fundef f = OK tf -> - funsig tf = funsig f. - Proof using . - unfold transl_fundef, transf_partial_fundef, transl_function; intros; - repeat destruct_match; crush; - match goal with H: OK _ = OK _ |- _ => inv H end; auto. - Qed. - Hint Resolve sig_transl_function : rtlgp. - - Hint Resolve Val.lessdef_same : rtlgp. - Hint Resolve regs_lessdef_regs : rtlgp. - - Lemma find_function_translated: - forall ros rs rs' f, - (forall x, rs !! x = rs' !! x) -> - find_function ge ros rs = Some f -> - exists tf, find_function tge ros rs' = Some tf - /\ transl_fundef f = OK tf. - Proof using TRANSL. - Ltac ffts := match goal with - | [ H: forall _, Val.lessdef _ _, r: Registers.reg |- _ ] => - specialize (H r); inv H - | [ H: Vundef = ?r, H1: Genv.find_funct _ ?r = Some _ |- _ ] => - rewrite <- H in H1 - | [ H: Genv.find_funct _ Vundef = Some _ |- _] => solve [inv H] - | _ => solve [exploit functions_translated; eauto] - end. - destruct ros; simplify; try rewrite <- H; - [| rewrite symbols_preserved; destruct_match; - try (apply function_ptr_translated); crush ]; - intros; - repeat ffts. - Qed. - - Lemma schedule_oracle_nil: - forall bb cfi, - schedule_oracle {| bb_body := nil; bb_exit := cfi |} bb = true -> - bb_body bb = nil /\ bb_exit bb = cfi. - Proof using . - unfold schedule_oracle, check_control_flow_instr. - simplify; repeat destruct_match; crush. - Qed. - - Lemma schedule_oracle_nil2: - forall cfi, - schedule_oracle {| bb_body := nil; bb_exit := cfi |} - {| bb_body := nil; bb_exit := cfi |} = true. - Proof using . - unfold schedule_oracle, check_control_flow_instr. - simplify; repeat destruct_match; crush. - Qed. - - Lemma eval_op_eq: - forall (sp0 : Values.val) (op : Op.operation) (vl : list Values.val) m, - Op.eval_operation ge sp0 op vl m = Op.eval_operation tge sp0 op vl m. - Proof using TRANSL. - intros. - destruct op; auto; unfold Op.eval_operation, Genv.symbol_address, Op.eval_addressing32; - [| destruct a; unfold Genv.symbol_address ]; - try rewrite symbols_preserved; auto. - Qed. - Hint Resolve eval_op_eq : rtlgp. - - Lemma eval_addressing_eq: - forall sp addr vl, - Op.eval_addressing ge sp addr vl = Op.eval_addressing tge sp addr vl. - Proof using TRANSL. - intros. - destruct addr; - unfold Op.eval_addressing, Op.eval_addressing32; - unfold Genv.symbol_address; - try rewrite symbols_preserved; auto. - Qed. - Hint Resolve eval_addressing_eq : rtlgp. - - Lemma ge_preserved_lem: - ge_preserved ge tge. - Proof using TRANSL. - unfold ge_preserved. - eauto with rtlgp. - Qed. - Hint Resolve ge_preserved_lem : rtlgp. - - Lemma lessdef_regmap_optget: - forall or rs rs', - regs_lessdef rs rs' -> - Val.lessdef (regmap_optget or Vundef rs) (regmap_optget or Vundef rs'). - Proof using. destruct or; crush. Qed. - Hint Resolve lessdef_regmap_optget : rtlgp. - - Lemma regmap_equiv_lessdef: - forall rs rs', - (forall x, rs !! x = rs' !! x) -> - regs_lessdef rs rs'. - Proof using. - intros; unfold regs_lessdef; intros. - rewrite H. apply Val.lessdef_refl. - Qed. - Hint Resolve regmap_equiv_lessdef : rtlgp. - - Lemma int_lessdef: - forall rs rs', - regs_lessdef rs rs' -> - (forall arg v, - rs !! arg = Vint v -> - rs' !! arg = Vint v). - Proof using. intros ? ? H; intros; specialize (H arg); inv H; crush. Qed. - Hint Resolve int_lessdef : rtlgp. - - Ltac semantics_simpl := - match goal with - | [ H: match_states _ _ |- _ ] => - let H2 := fresh "H" in - learn H as H2; inv H2 - | [ H: transl_function ?f = OK _ |- _ ] => - let H2 := fresh "TRANSL" in - learn H as H2; - unfold transl_function in H2; - destruct (check_scheduled_trees - (@fn_code RTLBlock.bb f) - (@fn_code RTLPar.bb (schedule f))) eqn:?; - [| discriminate ]; inv H2 - | [ H: context[check_scheduled_trees] |- _ ] => - let H2 := fresh "CHECK" in - learn H as H2; - eapply check_scheduled_trees_correct in H2; [| solve [eauto] ] - | [ H: schedule_oracle {| bb_body := nil; bb_exit := _ |} ?bb = true |- _ ] => - let H2 := fresh "SCHED" in - learn H as H2; - apply schedule_oracle_nil in H2 - | [ H: find_function _ _ _ = Some _, H2: forall x, ?rs !! x = ?rs' !! x |- _ ] => - learn H; exploit find_function_translated; try apply H2; eauto; inversion 1 - | [ H: Mem.free ?m _ _ _ = Some ?m', H2: Mem.extends ?m ?m'' |- _ ] => - learn H; exploit Mem.free_parallel_extends; eauto; intros - | [ H: Events.eval_builtin_args _ _ _ _ _ _, H2: regs_lessdef ?rs ?rs' |- _ ] => - let H3 := fresh "H" in - learn H; exploit Events.eval_builtin_args_lessdef; [apply H2 | | |]; - eauto with rtlgp; intro H3; learn H3 - | [ H: Events.external_call _ _ _ _ _ _ _ |- _ ] => - let H2 := fresh "H" in - learn H; exploit Events.external_call_mem_extends; - eauto; intro H2; learn H2 - | [ H: exists _, _ |- _ ] => inv H - | _ => progress simplify - end. - - Hint Resolve Events.eval_builtin_args_preserved : rtlgp. - Hint Resolve Events.external_call_symbols_preserved : rtlgp. - Hint Resolve set_res_lessdef : rtlgp. - Hint Resolve set_reg_lessdef : rtlgp. - Hint Resolve Op.eval_condition_lessdef : rtlgp. - - Hint Constructors Events.eval_builtin_arg: barg. - - Lemma eval_builtin_arg_eq: - forall A ge a v1 m1 e1 e2 sp, - (forall x, e1 x = e2 x) -> - @Events.eval_builtin_arg A ge e1 sp m1 a v1 -> - Events.eval_builtin_arg ge e2 sp m1 a v1. -Proof. induction 2; try rewrite H; eauto with barg. Qed. - - Lemma eval_builtin_args_eq: - forall A ge e1 sp m1 e2 al vl1, - (forall x, e1 x = e2 x) -> - @Events.eval_builtin_args A ge e1 sp m1 al vl1 -> - Events.eval_builtin_args ge e2 sp m1 al vl1. - Proof. - induction 2. - - econstructor; split. - - exploit eval_builtin_arg_eq; eauto. intros. - destruct IHlist_forall2 as [| y]. constructor; eauto. - constructor. constructor; auto. - constructor; eauto. - Qed. - - Lemma step_cf_instr_correct: - forall cfi t s s', - step_cf_instr ge s cfi t s' -> - forall r, - match_states s r -> - exists r', step_cf_instr tge r cfi t r' /\ match_states s' r'. - Proof using TRANSL. - induction 1; repeat semantics_simpl; - try solve [repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp)]. - { do 3 (try erewrite match_states_list by eauto; econstructor; eauto with rtlgp). - eapply eval_builtin_args_eq. eapply REG. - eapply Events.eval_builtin_args_preserved. eapply symbols_preserved. - eauto. - intros. - unfold regmap_setres. destruct res. - destruct (Pos.eq_dec x0 x); subst. - repeat rewrite Regmap.gss; auto. - repeat rewrite Regmap.gso; auto. - eapply REG. eapply REG. - } - { repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). - unfold regmap_optget. destruct or. rewrite REG. constructor; eauto. - constructor; eauto. - } - { exploit IHstep_cf_instr; eauto. simplify. - repeat (try erewrite match_states_list; eauto; econstructor; eauto with rtlgp). - erewrite eval_predf_pr_equiv; eauto. - } - Qed. - - Theorem transl_step_correct : - forall (S1 : RTLBlock.state) t S2, - RTLBlock.step ge S1 t S2 -> - forall (R1 : RTLPar.state), - match_states S1 R1 -> - exists R2, Smallstep.plus RTLPar.step tge R1 t R2 /\ match_states S2 R2. - Proof. - - induction 1; repeat semantics_simpl. - - { destruct bb; destruct x. - assert (bb_exit = bb_exit0). - { unfold schedule_oracle in *. simplify. - unfold check_control_flow_instr in *. - destruct_match; crush. - } - subst. - - exploit abstract_execution_correct; try eassumption. eapply ge_preserved_lem. - econstructor; eauto. - simplify. destruct x. inv H7. - - exploit step_cf_instr_correct; try eassumption. econstructor; eauto. - simplify. - - econstructor. econstructor. eapply Smallstep.plus_one. econstructor. - 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. - 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. } - { inv STACKS. inv H2. repeat econstructor; eauto. - intros. apply PTree_matches; eauto. } - Qed. - - Lemma transl_initial_states: - forall S, - RTLBlock.initial_state prog S -> - exists R, RTLPar.initial_state tprog R /\ match_states S R. - Proof. - induction 1. - exploit function_ptr_translated; eauto. intros [tf [A B]]. - econstructor; split. - econstructor. apply (Genv.init_mem_transf_partial TRANSL); eauto. - replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved; eauto. - symmetry; eapply match_program_main; eauto. - eexact A. - rewrite <- H2. apply sig_transl_function; auto. - constructor. auto. constructor. - Qed. - - Lemma transl_final_states: - forall S R r, - match_states S R -> RTLBlock.final_state S r -> RTLPar.final_state R r. - Proof. - intros. inv H0. inv H. inv STACKS. constructor. - Qed. - - Theorem transf_program_correct: - Smallstep.forward_simulation (RTLBlock.semantics prog) (RTLPar.semantics tprog). - Proof. - eapply Smallstep.forward_simulation_plus. - apply senv_preserved. - eexact transl_initial_states. - eexact transl_final_states. - exact transl_step_correct. - Qed. - -End CORRECTNESS. -#+end_src - -* License - -#+name: license -#+begin_src coq :tangle no -(* - * 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 . - *) -#+end_src diff --git a/src/Compiler.v b/src/Compiler.v index 8882686..cd05aa9 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -16,15 +16,21 @@ * along with this program. If not, see . *) -(** * Compiler Proof +(*| +============== +Compiler Proof +============== -This is the top-level module of the correctness proof and proves the final backwards simulation -correct. +This is the top-level module of the correctness proof and proves the final +backwards simulation correct. -** Imports +Imports +======= -We first need to import all of the modules that are used in the correctness proof, which includes -all of the passes that are performed in Vericert and the CompCert front end. *) +We first need to import all of the modules that are used in the correctness +proof, which includes all of the passes that are performed in Vericert and the +CompCert front end. +|*) Require compcert.backend.Selection. Require compcert.backend.RTL. @@ -65,10 +71,14 @@ Require vericert.hls.Memorygen. Require Import vericert.hls.HTLgenproof. -(** ** Declarations +(*| +Declarations +============ -We then need to declare the external OCaml functions used to print out intermediate steps in the -compilation, such as [print_RTL], [print_HTL] and [print_RTLBlock]. *) +We then need to declare the external OCaml functions used to print out +intermediate steps in the compilation, such as ``print_RTL``, ``print_HTL`` and +``print_RTLBlock``. +|*) Parameter print_RTL: Z -> RTL.program -> unit. Parameter print_HTL: Z -> HTL.program -> unit. @@ -86,16 +96,20 @@ Proof. intros; unfold print. destruct (printer prog); auto. Qed. -(** We also declare some new notation, which is also used in CompCert to combine the monadic results -of each pass. *) +(*| +We also declare some new notation, which is also used in CompCert to combine the +monadic results of each pass. +|*) Notation "a @@@ b" := (Compiler.apply_partial _ _ a b) (at level 50, left associativity). Notation "a @@ b" := (Compiler.apply_total _ _ a b) (at level 50, left associativity). -(** As printing is used in the translation but does not change the output, we need to prove that it -has no effect so that it can be removed during the proof. *) +(*| +As printing is used in the translation but does not change the output, we need +to prove that it has no effect so that it can be removed during the proof. +|*) Lemma compose_print_identity: forall (A: Type) (x: res A) (f: A -> unit), @@ -104,8 +118,10 @@ Proof. intros. destruct x; simpl. rewrite print_identity. auto. auto. Qed. -(** Finally, some optimisation passes are only activated by a flag, which is handled by the following -functions for partial and total passes. *) +(*| +Finally, some optimisation passes are only activated by a flag, which is handled +by the following functions for partial and total passes. +|*) Definition total_if {A: Type} (flag: unit -> bool) (f: A -> A) (prog: A) : A := @@ -156,11 +172,15 @@ Proof. intros. unfold match_if in *. destruct (flag tt). eauto. subst. apply forward_simulation_identity. Qed. -(** *** Top-level Translation +(*| +Top-level Translation +--------------------- -An optimised transformation function from [RTL] to [Verilog] can then be defined, which applies -the front end compiler optimisations of CompCert to the RTL that is generated and then performs the -two Vericert passes from RTL to HTL and then from HTL to Verilog. *) +An optimised transformation function from ``RTL`` to ``Verilog`` can then be +defined, which applies the front end compiler optimisations of CompCert to the +RTL that is generated and then performs the two Vericert passes from RTL to HTL +and then from HTL to Verilog. +|*) Definition transf_backend (r : RTL.program) : res Verilog.program := OK r @@ -184,8 +204,10 @@ Definition transf_backend (r : RTL.program) : res Verilog.program := @@ print (print_HTL 1)*) @@ Veriloggen.transl_program. -(** The transformation functions from RTL to Verilog are then added to the backend of the CompCert -transformations from Clight to RTL. *) +(*| +The transformation functions from RTL to Verilog are then added to the backend +of the CompCert transformations from Clight to RTL. +|*) Definition transf_hls (p : Csyntax.program) : res Verilog.program := OK p @@ -234,10 +256,13 @@ Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := @@ print (print_HTL 0) @@ Veriloggen.transl_program. -(** ** Correctness Proof +(*| +Correctness Proof +================= -Finally, the top-level definition for all the passes is defined, which combines the [match_prog] -predicates of each translation pass from C until Verilog. *) +Finally, the top-level definition for all the passes is defined, which combines +the ``match_prog`` predicates of each translation pass from C until Verilog. +|*) Local Open Scope linking_scope. @@ -260,14 +285,18 @@ Definition CompCert's_passes := ::: mkpass Veriloggenproof.match_prog ::: pass_nil _. -(** These passes are then composed into a larger, top-level [match_prog] predicate which matches a -C program directly with a Verilog program. *) +(*| +These passes are then composed into a larger, top-level ``match_prog`` predicate +which matches a C program directly with a Verilog program. +|*) Definition match_prog: Csyntax.program -> Verilog.program -> Prop := pass_match (compose_passes CompCert's_passes). -(** We then need to prove that this predicate holds, assuming that the translation is performed -using the [transf_hls] function declared above. *) +(*| +We then need to prove that this predicate holds, assuming that the translation +is performed using the ``transf_hls`` function declared above. +|*) Theorem transf_hls_match: forall p tp, @@ -369,12 +398,16 @@ Ltac DestructM := apply Verilog.semantics_determinate. Qed. -(** *** Backward Simulation +(*| +Backward Simulation +------------------- -The following theorem is a *backward simulation* between the C and Verilog, which proves the -semantics preservation of the translation. We can assume that the C and Verilog programs match, as -we have proven previously in [transf_hls_match] that our translation implies that the -[match_prog] predicate will hold. *) +The following theorem is a *backward simulation* between the C and Verilog, +which proves the semantics preservation of the translation. We can assume that +the C and Verilog programs match, as we have proven previously in +``transf_hls_match`` that our translation implies that the ``match_prog`` +predicate will hold. +|*) Theorem c_semantic_preservation: forall p tp, @@ -391,8 +424,11 @@ Proof. exact (proj2 (cstrategy_semantic_preservation _ _ H)). Qed. -(** We can then use [transf_hls_match] to prove the backward simulation where the assumption is -that the translation is performed using the [transf_hls] function and that it succeeds. *) +(*| +We can then use ``transf_hls_match`` to prove the backward simulation where the +assumption is that the translation is performed using the ``transf_hls`` +function and that it succeeds. +|*) Theorem transf_c_program_correct: forall p tp, @@ -402,9 +438,12 @@ Proof. intros. apply c_semantic_preservation. apply transf_hls_match; auto. Qed. -(** The final theorem of the semantic preservation of the translation of separate translation units -can also be proven correct, however, this is only because the translation fails if more than one -translation unit is passed to Vericert at the moment. *) +(*| +The final theorem of the semantic preservation of the translation of separate +translation units can also be proven correct, however, this is only because the +translation fails if more than one translation unit is passed to Vericert at the +moment. +|*) Theorem separate_transf_c_program_correct: forall c_units verilog_units c_program, diff --git a/src/bourdoncle/Bourdoncle.v b/src/bourdoncle/Bourdoncle.v index 9d470b7..ef1798a 100644 --- a/src/bourdoncle/Bourdoncle.v +++ b/src/bourdoncle/Bourdoncle.v @@ -1,4 +1,6 @@ -(** Type of a Bourdoncle component. *) +(*| +Type of a Bourdoncle component. +|*) Require Import List. Require Import BinPos. diff --git a/src/common/Maps.v b/src/common/Maps.v index 13ca276..cb9dc2c 100644 --- a/src/common/Maps.v +++ b/src/common/Maps.v @@ -29,8 +29,9 @@ Import PTree. Local Open Scope error_monad_scope. -(** Instance of traverse for [PTree] and [Errors]. This should maybe be generalised - in the future. *) +(*| +Instance of traverse for ``PTree`` and ``Errors``. This should maybe be generalised in the future. +|*) Module PTree. 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 a5b8a41ef22618c69db62dfeb71d7f38bbba34e2 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 26 Mar 2022 15:50:53 +0000 Subject: Add more files to .gitignore --- .gitignore | 3 +++ 1 file changed, 3 insertions(+) diff --git a/.gitignore b/.gitignore index 586f3a7..6568d8d 100644 --- a/.gitignore +++ b/.gitignore @@ -86,6 +86,7 @@ creduce_bug_*/ /docs/man/ /docs/manual/ /docs/src/ +/html/ *~ *.man @@ -94,3 +95,5 @@ creduce_bug_*/ /bin /share /results* + +_CoqProject -- 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 --- .gitignore | 4 +- Makefile | 3 +- doc/Makefile | 24 + doc/Makefile.extr | 39 ++ doc/README.org | 7 + doc/_static/css/custom.css | 11 + doc/_static/images/toolflow.pdf | Bin 0 -> 20490 bytes doc/_static/images/toolflow.png | Bin 0 -> 26864 bytes doc/_static/images/toolflow.svg | 1250 +++++++++++++++++++++++++++++++++++++++ doc/common.org | 15 + doc/conf.py | 85 +++ doc/documentation.org | 562 ++++++++++++++++++ doc/docutils.conf | 2 + doc/index.rst | 26 + doc/make.bat | 35 ++ doc/man.org | 89 +++ doc/res/coqdoc.css | 867 +++++++++++++++++++++++++++ doc/res/fdl.org | 489 +++++++++++++++ doc/res/install-deps.el | 11 + doc/res/publish.el | 23 + doc/vericert.rst | 530 +++++++++++++++++ docs/Makefile | 39 -- docs/README.org | 7 - docs/common.org | 15 - docs/documentation.org | 562 ------------------ docs/images/toolflow.pdf | Bin 20490 -> 0 bytes docs/images/toolflow.png | Bin 26864 -> 0 bytes docs/images/toolflow.svg | 1250 --------------------------------------- docs/man.org | 89 --- docs/res/coqdoc.css | 867 --------------------------- docs/res/fdl.org | 489 --------------- docs/res/install-deps.el | 11 - docs/res/publish.el | 23 - src/Compiler.v | 172 ++++-- src/hls/RTLBlockInstr.v | 55 +- 35 files changed, 4206 insertions(+), 3445 deletions(-) create mode 100644 doc/Makefile create mode 100644 doc/Makefile.extr create mode 100644 doc/README.org create mode 100644 doc/_static/css/custom.css create mode 100644 doc/_static/images/toolflow.pdf create mode 100644 doc/_static/images/toolflow.png create mode 100644 doc/_static/images/toolflow.svg create mode 100644 doc/common.org create mode 100644 doc/conf.py create mode 100644 doc/documentation.org create mode 100644 doc/docutils.conf create mode 100644 doc/index.rst create mode 100644 doc/make.bat create mode 100644 doc/man.org create mode 100644 doc/res/coqdoc.css create mode 100644 doc/res/fdl.org create mode 100644 doc/res/install-deps.el create mode 100644 doc/res/publish.el create mode 100644 doc/vericert.rst delete mode 100644 docs/Makefile delete mode 100644 docs/README.org delete mode 100644 docs/common.org delete mode 100644 docs/documentation.org delete mode 100644 docs/images/toolflow.pdf delete mode 100644 docs/images/toolflow.png delete mode 100644 docs/images/toolflow.svg delete mode 100644 docs/man.org delete mode 100644 docs/res/coqdoc.css delete mode 100644 docs/res/fdl.org delete mode 100644 docs/res/install-deps.el delete mode 100644 docs/res/publish.el diff --git a/.gitignore b/.gitignore index 6568d8d..5574171 100644 --- a/.gitignore +++ b/.gitignore @@ -83,9 +83,7 @@ obj_dir/ creduce_bug_*/ -/docs/man/ -/docs/manual/ -/docs/src/ +/doc/src/ /html/ *~ diff --git a/Makefile b/Makefile index a6aa914..3128c23 100644 --- a/Makefile +++ b/Makefile @@ -17,7 +17,8 @@ COQEXEC := $(COQBIN)coqtop $(COQINCLUDES) -batch -load-vernac-source COQMAKE := $(COQBIN)coq_makefile COQDOCFLAGS := --no-lib-name -l -ALECTRYON_OPTS := --html-minification --long-line-threshold 80 --coq-driver sertop_noexec $(COQINCLUDES) +ALECTRYON_OPTS := --no-header --html-minification --long-line-threshold 80 \ + --coq-driver sertop_noexec $(COQINCLUDES) VS := src/Compiler.v src/Simulator.v src/HLSOpts.v $(foreach d, common hls bourdoncle, $(wildcard src/$(d)/*.v)) LIT := docs/basic-block-generation.org docs/scheduler.org docs/scheduler-languages.org diff --git a/doc/Makefile b/doc/Makefile new file mode 100644 index 0000000..5662645 --- /dev/null +++ b/doc/Makefile @@ -0,0 +1,24 @@ +# Minimal makefile for Sphinx documentation +# + +# You can set these variables from the command line, and also +# from the environment for the first two. +SPHINXOPTS ?= +SPHINXBUILD ?= sphinx-build +SOURCEDIR = . +BUILDDIR = _build +SOURCE_DATE_EPOCH = $(shell git log -1 --format=%ct) + +VS_DOCS := ../src/Compiler.v ../src/hls/RTLBlockInstr.v + +# Put it first so that "make" without argument is like "make help". +help: + @$(SPHINXBUILD) -M help "$(SOURCEDIR)" "$(BUILDDIR)" $(SPHINXOPTS) $(O) + +.PHONY: help Makefile + +# Catch-all target: route all unknown targets to Sphinx using the new +# "make mode" option. $(O) is meant as a shortcut for $(SPHINXOPTS). +%: Makefile + $(foreach d,$(VS_DOCS),cp $(d) $(patsubst ../%,%,$(d));) + @$(SPHINXBUILD) -M $@ "$(SOURCEDIR)" "$(BUILDDIR)" $(SPHINXOPTS) $(O) diff --git a/doc/Makefile.extr b/doc/Makefile.extr new file mode 100644 index 0000000..9d4f361 --- /dev/null +++ b/doc/Makefile.extr @@ -0,0 +1,39 @@ +all: manual src man-html + +install-deps: + emacs --batch --load ./res/install-deps.el + +%.man: %.org + emacs --batch --file $< --load ./res/publish.el --funcall org-man-export-to-man + +%.html: %.org + emacs --batch --file $< --load ./res/publish.el --funcall org-html-export-to-html + +manual: + mkdir -p manual + emacs --batch --file documentation.org --load ./res/publish.el --funcall org-texinfo-export-to-texinfo + makeinfo --html --number-sections --no-split \ + --css-ref "https://www.gnu.org/software/emacs/manual.css" \ + vericert.texi -o ./manual/index.html + cp -r images ./manual/. + +man-html: man.html + mkdir -p man + cp man.html ./man/vericert.1.html + +vericert.1: man.man + cp $< $@ + +src: + $(MAKE) -C .. doc + cp -r ../html src + +upload: + tar caf docs.tar.xz manual man src + rsync docs.tar.xz "zk@leika.ymhg.org:~" + +.PHONY: all upload manual man src install-deps man-html + +clean: + rm -rf manual man src + rm -f docs.tar.xz diff --git a/doc/README.org b/doc/README.org new file mode 100644 index 0000000..4113ed9 --- /dev/null +++ b/doc/README.org @@ -0,0 +1,7 @@ +#+title: Vericert Documentation + +The documentation for [[https://github.com/ymherklotz/vericert][Vericert]], which is written in the [[/documentation.org][documentation.org]] file. + +* How to develop on documentation + +The documentation uses =hugo= to generate the website, diff --git a/doc/_static/css/custom.css b/doc/_static/css/custom.css new file mode 100644 index 0000000..66a1e35 --- /dev/null +++ b/doc/_static/css/custom.css @@ -0,0 +1,11 @@ +.alectryon-coqdoc .doc .code, +.alectryon-coqdoc .doc .comment, +.alectryon-coqdoc .doc .inlinecode, +.alectryon-mref, +.alectryon-block, .alectryon-io, +.alectryon-toggle-label, .alectryon-banner, pre, tt, code { + font-family: 'Iosevka Fixed Slab', 'Iosevka Slab Web', 'Iosevka Web', 'Iosevka Slab', + 'Iosevka', 'Fira Code', monospace; + font-feature-settings: "COQX" 1 /* Coq ligatures */, "XV00" 1 /* Legacy */, "calt" 1 /* Fallback */; + line-height: initial; +} diff --git a/doc/_static/images/toolflow.pdf b/doc/_static/images/toolflow.pdf new file mode 100644 index 0000000..5fcee67 Binary files /dev/null and b/doc/_static/images/toolflow.pdf differ diff --git a/doc/_static/images/toolflow.png b/doc/_static/images/toolflow.png new file mode 100644 index 0000000..601d6be Binary files /dev/null and b/doc/_static/images/toolflow.png differ diff --git a/doc/_static/images/toolflow.svg b/doc/_static/images/toolflow.svg new file mode 100644 index 0000000..0d8f39f --- /dev/null +++ b/doc/_static/images/toolflow.svg @@ -0,0 +1,1250 @@ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + diff --git a/doc/common.org b/doc/common.org new file mode 100644 index 0000000..aa27264 --- /dev/null +++ b/doc/common.org @@ -0,0 +1,15 @@ +#+author: Yann Herklotz +#+email: git@ymhg.org + +#+macro: version 1.2.2 +#+macro: modified 2022-02-24 + +#+macro: base_url https://vericert.ymhg.org + +#+macro: link src_emacs-lisp[:results raw]{(ymhg/link "$1" "$2")} + +#+macro: texinfo_head (eval (if (eq org-export-current-backend 'texinfo) "#+exclude_tags: noexport_texinfo" "#+exclude_tags: onlyexport_texinfo")) +#+macro: latex_head (eval (if (eq org-export-current-backend 'latex) "#+exclude_tags: noexport_latex" "#+exclude_tags: onlyexport_latex")) +#+macro: hugo_head (eval (if (eq org-export-current-backend 'hugo) "#+exclude_tags: noexport_hugo" "#+exclude_tags: onlyexport_hugo")) +{{{texinfo_head}}} +{{{latex_head}}} diff --git a/doc/conf.py b/doc/conf.py new file mode 100644 index 0000000..015f030 --- /dev/null +++ b/doc/conf.py @@ -0,0 +1,85 @@ +# Configuration file for the Sphinx documentation builder. +# +# This file only contains a selection of the most common options. For a full +# list see the documentation: +# https://www.sphinx-doc.org/en/master/usage/configuration.html + +# -- Path setup -------------------------------------------------------------- + +# If extensions (or modules to document with autodoc) are in another directory, +# add these directories to sys.path here. If the directory is relative to the +# documentation root, use os.path.abspath to make it absolute, like shown here. +# +# import os +# import sys +# sys.path.insert(0, os.path.abspath('.')) + + +# -- Project information ----------------------------------------------------- + +project = 'Vericert' +copyright = '2022 Yann Herklotz, John Wickerson' +author = 'Yann Herklotz, John Wickerson' + +# The full version, including alpha/beta/rc tags +release = 'v1.2.2' + + +# -- General configuration --------------------------------------------------- + +# Add any Sphinx extension module names here, as strings. They can be +# extensions coming with Sphinx (named 'sphinx.ext.*') or your custom +# ones. +extensions = [ "alectryon.sphinx" ] + +# Add any paths that contain templates here, relative to this directory. +templates_path = ['_templates'] + +# List of patterns, relative to source directory, that match files and +# directories to ignore when looking for source files. +# This pattern also affects html_static_path and html_extra_path. +exclude_patterns = ['_build', 'Thumbs.db', '.DS_Store'] + +pygments_style = "emacs" + + +# -- Options for HTML output ------------------------------------------------- + +# The theme to use for HTML and HTML Help pages. See the documentation for +# a list of builtin themes. +# +html_theme = 'alabaster' + +# Add any paths that contain custom static files (such as style sheets) here, +# relative to this directory. They are copied after the builtin static files, +# so a file named "default.css" will overwrite the builtin "default.css". +html_static_path = ['_static'] + +html_css_files = [ + 'css/custom.css', +] + + +# -- LaTeX configuration ----------------------------------------------------- + +latex_engine = "xelatex" + + +# -- Alectryon configuration ------------------------------------------------- + +import alectryon.docutils +alectryon.docutils.CACHE_DIRECTORY = "_build/alectryon/" +alectryon.docutils.LONG_LINE_THRESHOLD = 80 + +alectryon.docutils.AlectryonTransform.DRIVER_ARGS["sertop"] = [ + "-R" "../src,vericert", + "-R" "../lib/CompCert/lib,compcert.lib", + "-R" "../lib/CompCert/common,compcert.common", + "-R" "../lib/CompCert/verilog,compcert.verilog", + "-R" "../lib/CompCert/backend,compcert.backend", + "-R" "../lib/CompCert/cfrontend,compcert.cfrontend", + "-R" "../lib/CompCert/driver,compcert.driver", + "-R" "../lib/CompCert/cparser,compcert.cparser", + "-R" "../lib/CompCert/flocq,Flocq", + "-R" "../lib/CompCert/MenhirLib,MenhirLib" +] diff --git a/doc/documentation.org b/doc/documentation.org new file mode 100644 index 0000000..85bc8fd --- /dev/null +++ b/doc/documentation.org @@ -0,0 +1,562 @@ +#+title: Vericert Manual +#+subtitle: Release {{{version}}} +#+author: Yann Herklotz +#+email: git@ymhg.org +#+language: en + +* Introduction +:PROPERTIES: +:EXPORT_FILE_NAME: _index +:EXPORT_HUGO_SECTION: docs +:CUSTOM_ID: docs +:END: + +Vericert translates C code into a hardware description language called Verilog, which can then be +synthesised into hardware, to be placed onto a field-programmable gate array (FPGA) or +application-specific integrated circuit (ASIC). + +#+attr_html: :width "600px" +#+caption: Current design of Vericert, where HTL is an intermediate language representing a finite state machine with data-path (FSMD) and Verilog is the target language. +#+name: fig:design +[[./images/toolflow.png]] + +The design shown in Figure [[fig:design]] shows how Vericert leverages an existing verified C compiler +called [[https://compcert.org/compcert-C.html][CompCert]] to perform this translation. + +#+texinfo: @insertcopying + +* COPYING +:PROPERTIES: +:COPYING: t +:END: + +Copyright (C) 2019-2022 Yann Herklotz. + +#+begin_quote +Permission is granted to copy, distribute and/or modify this document +under the terms of the GNU Free Documentation License, Version 1.3 +or any later version published by the Free Software Foundation; +with no Invariant Sections, no Front-Cover Texts, and no Back-Cover +Texts. A copy of the license is included in the section entitled ``GNU +Free Documentation License''. +#+end_quote + +* Building Vericert +:PROPERTIES: +:EXPORT_FILE_NAME: building +:EXPORT_HUGO_SECTION: docs +:CUSTOM_ID: building +:END: + +#+transclude: [[file:~/projects/vericert/README.org::#building][file:../README.org::#building]] :only-contents :exclude-elements "headline property-drawer" +#+transclude: [[file:~/projects/vericert/README.org::#downloading-compcert][file:../README.org::#downloading-compcert]] :level 2 +#+transclude: [[file:~/projects/vericert/README.org::#setting-up-nix][file:../README.org::#setting-up-nix]] :level 2 +#+transclude: [[file:~/projects/vericert/README.org::#makefile-build][file:../README.org::#makefile-build]] :level 2 + +** Testing + +To test out ~vericert~ you can try the following examples which are in the test folder using the +following: + +#+begin_src shell +./bin/vericert test/loop.c -o loop.v +./bin/vericert test/conditional.c -o conditional.v +./bin/vericert test/add.c -o add.v +#+end_src + +Or by running the test suite using the following command: + +#+begin_src shell +make test +#+end_src + +* Using Vericert +:PROPERTIES: +:CUSTOM_ID: using-vericert +:END: + +Vericert can be used to translate a subset of C into Verilog. As a simple example, consider the +following C file (~main.c~): + +#+begin_src C +void matrix_multiply(int first[2][2], int second[2][2], int multiply[2][2]) { + int sum = 0; + for (int c = 0; c < 2; c++) { + for (int d = 0; d < 2; d++) { + for (int k = 0; k < 2; k++) { + sum = sum + first[c][k]*second[k][d]; + } + multiply[c][d] = sum; + sum = 0; + } + } +} + +int main() { + int f[2][2] = {{1, 2}, {3, 4}}; + int s[2][2] = {{5, 6}, {7, 8}}; + int m[2][2] = {{0, 0}, {0, 0}}; + + matrix_multiply(f, s, m); + return m[1][1]; +} +#+end_src + +It can be compiled using the following command, assuming that vericert is somewhere on the path. + +#+begin_src shell +vericert main.c -o main.v +#+end_src + +The Verilog file contains a top-level test-bench, which can be given to any Verilog simulator to +simulate the hardware, which should give the same result as executing the C code. Using [[http://iverilog.icarus.com/][Icarus +Verilog]] as an example: + +#+begin_src shell +iverilog -o main_v main.v +#+end_src + +When executing, it should therefore print the following: + +#+begin_src shell +$ ./main_v +finished: 50 +#+end_src + +This gives the same result as executing the C in the following way: + +#+begin_src shell +$ gcc -o main_c main.c +$ ./main_c +$ echo $? +50 +#+end_src + +** Man pages + +#+transclude: [[file:man.org][file:man.org]] :exclude-elements "keyword" :level 3 + +* Unreleased Features +:PROPERTIES: +:EXPORT_FILE_NAME: unreleased +:EXPORT_HUGO_SECTION: docs +:CUSTOM_ID: unreleased-features +:END: + +The following are unreleased features in Vericert that are currently being worked on and have not +been completely proven correct yet. Currently this includes features such as: + +- [[#scheduling][scheduling]], +- [[#operation-chaining][operation chaining]], +- [[#if-conversion][if-conversion]], and +- [[#functions][functions]]. + +This page gives some preliminary information on how the features are implemented and how the proofs +for the features are being done. Once these features are properly implemented, they will be added +to the proper documentation. + +** Scheduling +:PROPERTIES: +:CUSTOM_ID: scheduling +:END: +#+cindex: scheduling + +Scheduling is an optimisation which is used to run various instructions in parallel that are +independent to each other. + +** Operation Chaining +:PROPERTIES: +:CUSTOM_ID: operation-chaining +:END: + +Operation chaining is an optimisation that can be added on to scheduling and allows for the +sequential execution of instructions in a clock cycle, while executing other instructions in +parallel in the same clock cycle. + +** If-conversion +:PROPERTIES: +:CUSTOM_ID: if-conversion +:END: + +If-conversion is an optimisation which can turn code with simple control flow into a single block +(called a hyper-block), using predicated instructions. + +** Functions +:PROPERTIES: +:CUSTOM_ID: functions +:END: + +Functions are currently only inlined in Vericert, however, we are working on a proper interface to +integrate function calls into the hardware. + +* Coq Style Guide + :PROPERTIES: + :CUSTOM_ID: coq-style-guide + :EXPORT_FILE_NAME: coq-style-guide + :EXPORT_HUGO_SECTION: docs + :END: + +This style guide was taken from [[https://github.com/project-oak/silveroak][Silveroak]], it outlines code style for Coq code in this +repository. There are certainly other valid strategies and opinions on Coq code style; this is laid +out purely in the name of consistency. For a visual example of the style, see the [[#example][example]] at the +bottom of this file. + +** Code organization + :PROPERTIES: + :CUSTOM_ID: code-organization + :END: +*** Legal banner + :PROPERTIES: + :CUSTOM_ID: legal-banner + :END: + +- Files should begin with a copyright/license banner, as shown in the example above. + +*** Import statements + :PROPERTIES: + :CUSTOM_ID: import-statements + :END: + +- =Require Import= statements should all go at the top of the file, followed by file-wide =Import= + statements. + + - =Import=s often contain notations or typeclass instances that might override notations or + instances from another library, so it's nice to highlight them separately. + +- One =Require Import= statement per line; it's easier to scan that way. +- =Require Import= statements should use "fully-qualified" names (e.g. =Require Import + Coq.ZArith.ZArith= instead of =Require Import ZArith=). + + - Use the =Locate= command to find the fully-qualified name! + +- =Require Import='s should go in the following order: + + 1. Standard library dependencies (start with =Coq.=) + 2. External dependencies (anything outside the current project) + 3. Same-project dependencies + +- =Require Import='s with the same root library (the name before the first =.=) should be grouped + together. Within each root-library group, they should be in alphabetical order (so =Coq.Lists.List= + before =Coq.ZArith.ZArith=). + +*** Notations and scopes + :PROPERTIES: + :CUSTOM_ID: notations-and-scopes + :END: + +- Any file-wide =Local Open Scope='s should come immediately after the =Import=s (see example). + + - Always use =Local Open Scope=; just =Open Scope= will sneakily open the scope for those who import + your file. + +- Put notations in their own separate modules or files, so that those who import your file can + choose whether or not they want the notations. + + - Conflicting notations can cause a lot of headache, so it comes in very handy to leave this + flexibility! + +** Formatting + :PROPERTIES: + :CUSTOM_ID: formatting + :END: +*** Line length + :PROPERTIES: + :CUSTOM_ID: line-length + :END: + +- Maximum line length 80 characters. + + - Many Coq IDE setups divide the screen in half vertically and use only half to display source + code, so more than 80 characters can be genuinely hard to read on a laptop. + +*** Whitespace and indentation + :PROPERTIES: + :CUSTOM_ID: whitespace-and-indentation + :END: + +- No trailing whitespace. + +- Spaces, not tabs. + +- Files should end with a newline. + + - Many editors do this automatically on save. + +- Colons may be either "English-spaced", with no space before the colon and one space after (=x: nat=) + or "French-spaced", with one space before and after (=x : nat=). + +- Default indentation is 2 spaces. + + - Keeping this small prevents complex proofs from being indented ridiculously far, and matches IDE + defaults. + +- Use 2-space indents if inserting a line break immediately after: + + - =Proof.= + - =fun <...> =>= + - =forall <...>,= + - =exists <....>,= + +- The style for indenting arguments in function application depends on where you make a line + break. If you make the line break immediately after the function name, use a 2-space + indent. However, if you make it after one or more arguments, align the next line with the first + argument: + + #+begin_src coq + (Z.pow + 1 2) + (Z.pow 1 2 3 + 4 5 6) + #+end_src + +- =Inductive= cases should not be indented. Example: + + #+begin_src coq + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo + . + #+end_src + +- =match= or =lazymatch= cases should line up with the "m" in =match= or "l" in =lazymatch=, as in the + following examples: + + #+begin_src coq + match x with + | 3 => true + | _ => false + end. + + lazymatch x with + | 3 => idtac + | _ => fail "Not equal to 3:" x + end. + + repeat match goal with + | _ => progress subst + | _ => reflexivity + end. + + do 2 lazymatch goal with + | |- context [eq] => idtac + end. + #+end_src + +** Definitions and Fixpoints + :PROPERTIES: + :CUSTOM_ID: definitions-and-fixpoints + :END: + +- It's okay to leave the return type of =Definition='s and =Fixpoint='s implicit (e.g. ~Definition x := 5~ + instead of ~Definition x : nat := 5~) when the type is very simple or obvious (for instance, the + definition is in a file which deals exclusively with operations on =Z=). + +** Inductives + :PROPERTIES: + :CUSTOM_ID: inductives + :END: + +- The =.= ending an =Inductive= can be either on the same line as the last case or on its own line + immediately below. That is, both of the following are acceptable: + + #+begin_src coq + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo + . + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo. + #+end_src + +** Lemma/Theorem statements + :PROPERTIES: + :CUSTOM_ID: lemmatheorem-statements + :END: + +- Generally, use =Theorem= for the most important, top-level facts you prove and =Lemma= for everything + else. +- Insert a line break after the colon in the lemma statement. +- Insert a line break after the comma for =forall= or =exist= quantifiers. +- Implication arrows (=->=) should share a line with the previous hypothesis, not the following one. +- There is no need to make a line break after every =->=; short preconditions may share a line. + +** Proofs and tactics + :PROPERTIES: + :CUSTOM_ID: proofs-and-tactics + :END: + +- Use the =Proof= command (lined up vertically with =Lemma= or =Theorem= it corresponds to) to open a + proof, and indent the first line after it 2 spaces. + +- Very small proofs (where =Proof. Qed.= is <= 80 characters) can go all in one line. + +- When ending a proof, align the ending statement (=Qed=, =Admitted=, etc.) with =Proof=. + +- Avoid referring to autogenerated names (e.g. =H0=, =n0=). It's okay to let Coq generate these names, + but you should not explicitly refer to them in your proof. So =intros; my_solver= is fine, but + =intros; apply H1; my_solver= is not fine. + + - You can force a non-autogenerated name by either putting the variable before the colon in the + lemma statement (=Lemma foo x : ...= instead of =Lemma foo : forall x, ...=), or by passing + arguments to =intros= (e.g. =intros ? x= to name the second argument =x=) + +- This way, the proof won't break when new hypotheses are added or autogenerated variable names + change. + +- Use curly braces ={}= for subgoals, instead of bullets. + +- /Never write tactics with more than one subgoal focused./ This can make the proof very confusing to + step through! If you have more than one subgoal, use curly braces. + +- Consider adding a comment after the opening curly brace that explains what case you're in (see + example). + + - This is not necessary for small subgoals but can help show the major lines of reasoning in large + proofs. + +- If invoking a tactic that is expected to return multiple subgoals, use =[ | ... | ]= before the =.= to + explicitly specify how many subgoals you expect. + + - Examples: =split; [ | ].= =induction z; [ | | ].= + - This helps make code more maintainable, because it fails immediately if your tactic no longer + solves as many subgoals as expected (or unexpectedly solves more). + +- If invoking a string of tactics (composed by =;=) that will break the goal into multiple subgoals + and then solve all but one, still use =[ ]= to enforce that all but one goal is solved. + + - Example: =split; try lia; [ ]=. + +- Tactics that consist only of =repeat=-ing a procedure (e.g. =repeat match=, =repeat first=) should + factor out a single step of that procedure a separate tactic called =_step=, because + the single-step version is much easier to debug. For instance: + + #+begin_src coq + Ltac crush_step := + match goal with + | _ => progress subst + | _ => reflexivity + end. + Ltac crush := repeat crush_step. + #+end_src + +** Naming + :PROPERTIES: + :CUSTOM_ID: naming + :END: + +- Helper proofs about standard library datatypes should go in a module that is named to match the + standard library module (see example). + + - This makes the helper proofs look like standard-library ones, which is helpful for categorizing + them if they're genuinely at the standard-library level of abstraction. + +- Names of modules should start with capital letters. +- Names of inductives and their constructors should start with capital letters. +- Names of other definitions/lemmas should be snake case. + +** Example + :PROPERTIES: + :CUSTOM_ID: example + :END: +A small standalone Coq file that exhibits many of the style points. + +#+begin_src coq +(* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2021 Name + * + * + *) + + Require Import Coq.Lists.List. + Require Import Coq.micromega.Lia. + Require Import Coq.ZArith.ZArith. + Import ListNotations. + Local Open Scope Z_scope. + + (* Helper proofs about standard library integers (Z) go within [Module Z] so + that they match standard-library Z lemmas when used. *) + Module Z. + Lemma pow_3_r x : x ^ 3 = x * x * x. + Proof. lia. Qed. (* very short proofs can go all on one line *) + + Lemma pow_4_r x : x ^ 4 = x * x * x * x. + Proof. + change 4 with (Z.succ (Z.succ (Z.succ (Z.succ 0)))). + repeat match goal with + | _ => rewrite Z.pow_1_r + | _ => rewrite Z.pow_succ_r by lia + | |- context [x * (?a * ?b)] => + replace (x * (a * b)) with (a * b * x) by lia + | _ => reflexivity + end. + Qed. + End Z. + (* Now we can access the lemmas above as Z.pow_3_r and Z.pow_4_r, as if they + were in the ZArith library! *) + + Definition bar (x y : Z) := x ^ (y + 1). + + (* example with a painfully manual proof to show case formatting *) + Lemma bar_upper_bound : + forall x y a, + 0 <= x <= a -> 0 <= y -> + 0 <= bar x y <= a ^ (y + 1). + Proof. + (* avoid referencing autogenerated names by explicitly naming variables *) + intros x y a Hx Hy. revert y Hy x a Hx. + (* explicitly indicate # subgoals with [ | ... | ] if > 1 *) + cbv [bar]; refine (natlike_ind _ _ _); [ | ]. + { (* y = 0 *) + intros; lia. } + { (* y = Z.succ _ *) + intros. + rewrite Z.add_succ_l, Z.pow_succ_r by lia. + split. + { (* 0 <= bar x y *) + apply Z.mul_nonneg_nonneg; [ lia | ]. + apply Z.pow_nonneg; lia. } + { (* bar x y < a ^ y *) + rewrite Z.pow_succ_r by lia. + apply Z.mul_le_mono_nonneg; try lia; + [ apply Z.pow_nonneg; lia | ]. + (* For more flexible proofs, use match statements to find hypotheses + rather than referring to them by autogenerated names like H0. In this + case, we'll take any hypothesis that applies to and then solves the + goal. *) + match goal with H : _ |- _ => apply H; solve [auto] end. } } + Qed. + + (* Put notations in a separate module or file so that importers can + decide whether or not to use them. *) + Module BarNotations. + Infix "#" := bar (at level 40) : Z_scope. + Notation "x '##'" := (bar x x) (at level 40) : Z_scope. + End BarNotations. +#+end_src + +* Index - Features +:PROPERTIES: +:CUSTOM_ID: cindex +:APPENDIX: t +:INDEX: cp +:DESCRIPTION: Key concepts & features +:END: + +* Export Setup :noexport: + +#+setupfile: common.org + +#+export_file_name: vericert.texi + +#+texinfo_dir_category: High-level synthesis tool +#+texinfo_dir_title: Vericert +#+texinfo_dir_desc: Formally verified high-level synthesis tool + +* GNU Free Documentation License +:PROPERTIES: +:appendix: t +:END: + +#+include: res/fdl.org diff --git a/doc/docutils.conf b/doc/docutils.conf new file mode 100644 index 0000000..1bf4d83 --- /dev/null +++ b/doc/docutils.conf @@ -0,0 +1,2 @@ +[restructuredtext parser] +syntax_highlight = short diff --git a/doc/index.rst b/doc/index.rst new file mode 100644 index 0000000..ebb99df --- /dev/null +++ b/doc/index.rst @@ -0,0 +1,26 @@ +.. Vericert documentation master file, created by + sphinx-quickstart on Sat Mar 26 18:15:40 2022. + You can adapt this file completely to your liking, but it should at least + contain the root `toctree` directive. + +Vericert's Documentation +======================== + +.. toctree:: + :maxdepth: 2 + :caption: Content: + + vericert + +.. toctree:: + :maxdepth: 1 + :caption: Sources: + + src/Compiler + src/hls/RTLBlockInstr + +Indices +================== + +* :ref:`genindex` +* :ref:`search` diff --git a/doc/make.bat b/doc/make.bat new file mode 100644 index 0000000..153be5e --- /dev/null +++ b/doc/make.bat @@ -0,0 +1,35 @@ +@ECHO OFF + +pushd %~dp0 + +REM Command file for Sphinx documentation + +if "%SPHINXBUILD%" == "" ( + set SPHINXBUILD=sphinx-build +) +set SOURCEDIR=. +set BUILDDIR=_build + +if "%1" == "" goto help + +%SPHINXBUILD% >NUL 2>NUL +if errorlevel 9009 ( + echo. + echo.The 'sphinx-build' command was not found. Make sure you have Sphinx + echo.installed, then set the SPHINXBUILD environment variable to point + echo.to the full path of the 'sphinx-build' executable. Alternatively you + echo.may add the Sphinx directory to PATH. + echo. + echo.If you don't have Sphinx installed, grab it from + echo.https://www.sphinx-doc.org/ + exit /b 1 +) + +%SPHINXBUILD% -M %1 %SOURCEDIR% %BUILDDIR% %SPHINXOPTS% %O% +goto end + +:help +%SPHINXBUILD% -M help %SOURCEDIR% %BUILDDIR% %SPHINXOPTS% %O% + +:end +popd diff --git a/doc/man.org b/doc/man.org new file mode 100644 index 0000000..cb6143f --- /dev/null +++ b/doc/man.org @@ -0,0 +1,89 @@ +#+title: vericert +#+man_class_options: :section-id "1" +#+options: toc:nil num:nil +#+html_head_extra: + +* NAME + +vericert - A formally verified high-level synthesis tool. + +* SYNOPSYS + +*vericert* [ *OPTION* ]... [ *FILE* ]... + +* DESCRIPTION + +** HLS Options: + +- --no-hls :: Do not use HLS and generate standard flow +- --simulate :: Simulate the result with the Verilog semantics +- --debug-hls :: Add debug logic to the Verilog +- --initialise-stack :: initialise the stack to all 0s + +** HLS Optimisations: + +- -fschedule :: Schedule the resulting hardware [off] +- -fif-conversion :: If-conversion optimisation [off] + +** General options: + +- -stdlib :: Set the path of the Compcert run-time library +- -v :: Print external commands before invoking them +- -timings :: Show the time spent in various compiler passes +- -version :: Print the version string and exit +- -target :: Generate code for the given target +- -conf :: Read configuration from file +- @ :: Read command line options from + +** Tracing Options: + +- -dprepro :: Save C file after preprocessing in .i +- -dparse :: Save C file after parsing and elaboration in .parsed.c +- -dc :: Save generated C in .compcert.c +- -dclight :: Save generated Clight in .light.c +- -dcminor :: Save generated Cminor in .cm +- -drtl :: Save RTL at various optimization points in .rtl. +- -drtlblock :: Save RTLBlock .rtlblock +- -dhtl :: Save HTL before Verilog generation .htl +- -dltl :: Save LTL after register allocation in .ltl +- -dmach :: Save generated Mach code in .mach +- -dasm :: Save generated assembly in .s +- -dall :: Save all generated intermediate files in . +- -sdump :: Save info for post-linking validation in .json +- -o :: Generate output in + +** Diagnostic options: + +- -Wall :: Enable all warnings +- -W :: Enable the specific +- -Wno- :: Disable the specific +- -Werror :: Make all warnings into errors +- -Werror= :: Turn into an error +- -Wno-error= :: Turn into a warning even if -Werror is specified +- -Wfatal-errors :: Turn all errors into fatal errors aborting the compilation +- -fdiagnostics-color :: Turn on colored diagnostics +- -fno-diagnostics-color :: Turn of colored diagnostics +- -fmax-errors= :: Maximum number of errors to report +- -fdiagnostics-show-option :: Print the option name with mappable diagnostics +- -fno-diagnostics-show-option :: Turn of printing of options with mappable diagnostics + +* AUTHOR + +Written by Yann Herklotz, Michalis Pardalos, James Pollard, Nadesh Ramanathan and John Wickerson. + +* COPYRIGHT + +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 . diff --git a/doc/res/coqdoc.css b/doc/res/coqdoc.css new file mode 100644 index 0000000..5572aaa --- /dev/null +++ b/doc/res/coqdoc.css @@ -0,0 +1,867 @@ +body { + padding: 0px 0px; + margin: 0px 0px; + padding-left: 1em; + background-color: white; + font-family: sans-serif; + background-repeat: no-repeat; + background-size: 100%; +} + +#page { + display: block; + padding: 0px; + margin: 0px; +} + +#header { + min-height: 100px; + max-width: 760px; + margin: 0 auto; + padding-left: 80px; + padding-right: 80px; + padding-top: 30px; +} + +#header h1 { + padding: 0; + margin: 0; +} + +/* Menu */ +ul#menu { + padding: 0; + display: block; + margin: auto; +} + +ul#menu li, div.button { + display: inline-block; + background-color: white; + padding: 5px; + font-size: .70em; + text-transform: uppercase; + width: 30%; + text-align: center; + font-weight: 600; +} + +div.button { + margin-top:5px; + width: 40%; +} + +#button_block {margin-top:50px;} + +ul#menu a.hover li { + background-color: red; +} + +/* Contents */ + +#main, #main_home { + display: block; + padding: 80px; + padding-top: 60px; /* BCP */ + font-size: 100%; + line-height: 100%; + max-width: 760px; + background-color: #ffffff; + margin: 0 auto; +} + +#main_home { + background-color: rgba(255, 255, 255, 0.5); +} + +#index_content div.intro p { + font-size: 12px; +} + +#main h1 { + /* line-height: 80%; */ + line-height: normal; + padding-top: 3px; + padding-bottom: 4px; + /* Demitri: font-size: 22pt; */ + font-size: 200%; /* BCP */ +} + +/* allow for multi-line headers */ +#main a.idref:visited {color : #416DFF; text-decoration : none; } +#main a.idref:link {color : #416DFF; text-decoration : none; } +#main a.idref:hover {text-decoration : none; } +#main a.idref:active {text-decoration : none; } + +#main a.modref:visited {color : #416DFF; text-decoration : none; } +#main a.modref:link {color : #416DFF; text-decoration : none; } +#main a.modref:hover {text-decoration : none; } +#main a.modref:active {text-decoration : none; } + +#main .keyword { color : #697f2f } +#main { color: black } + +/* General section class - applies to all .section IDs */ +.section { + padding-top: 12px; + padding-bottom: 11px; + padding-left: 8px; + margin-top: 5px; + margin-bottom: 3px; + margin-top: 18px; + font-size : 125%; + color: #ffffff; +} + +/* Book title in header */ +.booktitleinheader { + color: #000000; + text-transform: uppercase; + font-weight: 300; + letter-spacing: 1px; + font-size: 125%; + margin-left: 0px; + margin-bottom: 22px; + } + +/* Chapter titles */ +.libtitle { + max-width: 860px; + text-transform: uppercase; + margin: 5px auto; + font-weight: 500; + padding-bottom: 2px; + font-size: 120%; + letter-spacing: 3px; + } + +.subtitle { + display: block; + padding-top: 10px; + font-size: 70%; + line-height: normal; +} + +h2.section { + color: #2a2c71; + background-color: transparent; + padding-left: 0px; + padding-top: 0px; + padding-bottom: 0px; + /* margin-top: 0px; */ + margin-top: 9px; /* BCP 2/20 */ + font-size : 135%; } + +h3.section { + /* background-color: rgb(90%,90%,100%); */ + background-color: white; + /* padding-left: 8px; */ + padding-left: 0px; + /* padding-top: 7px; */ + padding-top: 0px; + /* padding-bottom: 0px; */ + padding-bottom: 0.5em; + /* margin-top: 9px; */ + /* margin-top: 4px; (BCP 2/20) */ + margin-top: 9px; /* BCP 2/20 */ + font-size : 115%; + color:black; +} + +h4.section { + margin-top: .2em; + background-color: white; + color: #2a2c71; + padding-left: 0px; + padding-top: 0px; + padding-bottom: 0.5em; /* 0px; */ + font-size : 100%; + font-style : bold; + text-decoration : underline; +} + +#header p { + font-size: 13px; +} + +/* Sets up Main ID and margins */ + +#main .doc { + margin: 0px auto; + font-size: 14px; + line-height: 22px; + /* max-width: 570px; */ + color: black; + /* text-align: justify; */ + border-style: plain; + /* This might work better after changing back to standard coqdoc: */ + padding-top: 10px; + /* padding-top: 18px; */ +} + +.quote { + margin-left: auto; + margin-right: auto; + width: 40em; + color: darkred; +} + +.loud { + color: darkred; +} + +pre { + margin-top: 0px; + margin-bottom: 0px; +} + +.inlinecode { + display: inline; + /* font-size: 125%; */ + color: #444444; + font-family: monospace } + +.doc .inlinecode { + display: inline; + font-size: 105%; + color: rgb(35%,35%,70%); + font-family: monospace } + +.doc .inlinecode .id { +/* I am changing this to white in style below: + color: rgb(35%,35%,70%); +*/ +} + +h1 .inlinecode .id, h1.section span.inlinecode { + color: #ffffff; +} + +.inlinecodenm { + display: inline; + /* font-size: 125%; */ + color: #444444; +} + +.doc .inlinecodenm { + display: inline; + color: rgb(35%,35%,70%); +} + +.doc .inlinecodenm .id { + color: rgb(35%,35%,70%); +} + +.doc .code { + display: inline; + font-size: 110%; + color: rgb(35%,35%,70%); + font-family: monospace; + padding-left: 0px; +} + +.comment { + display: inline; + font-family: monospace; + color: rgb(50%,50%,80%); +} + +.inlineref { + display: inline; + /* font-family: monospace; */ + color: rgb(50%,50%,80%); +} + +.show::before { + /* content: "more"; */ + content: "+"; +} + +.show { + background-color: rgb(93%,93%,93%); + display: inline; + font-family: monospace; + font-size: 60%; + padding-top: 1px; + padding-bottom: 2px; + padding-left: 4px; + padding-right: 4px; + color: rgb(60%,60%,60%); +} + +/* +.show { + display: inline; + font-family: monospace; + font-size: 60%; + padding-top: 0px; + padding-bottom: 0px; + padding-left: 10px; + border: 1px; + border-style: solid; + color: rgb(75%,75%,85%); +} +*/ + +.proofbox { + font-size: 90%; + color: rgb(75%,75%,75%); +} + +#main .less-space { + margin-top: -12px; +} + +/* Inline quizzes */ +.quiz:before { + color: rgb(40%,40%,40%); + /* content: "- Quick Check -" ; */ + display: block; + text-align: center; + margin-bottom: 5px; +} +.quiz { + border: 4px; + border-color: rgb(80%,80%,80%); + /* + margin-left: 40px; + margin-right: 100px; + */ + padding: 5px; + padding-left: 8px; + padding-right: 8px; + margin-top: 10px; + margin-bottom: 15px; + border-style: solid; +} + +/* For textual ones... */ +.show-old { + display: inline; + font-family: monospace; + font-size: 80%; + padding-top: 0px; + padding-bottom: 0px; + padding-left: 3px; + padding-right: 3px; + border: 1px; + margin-top: 5px; /* doesn't work?! */ + border-style: solid; + color: rgb(75%,75%,85%); +} + +.largebr { + margin-top: 10px; +} + +.code { + padding-left: 20px; + font-size: 14px; + font-family: monospace; + line-height: 17px; + margin-top: 9px; +} + +/* Working: +.code { + display: block; + padding-left: 0px; + font-size: 110%; + font-family: monospace; + } +*/ + +.code-space { + max-width: 50em; + margin-top: 0em; + /* margin-bottom: -.5em; */ + margin-left: auto; + margin-right: auto; +} + +.code-tight { + max-width: 50em; + margin-top: .6em; + /* margin-bottom: -.7em; */ + margin-left: auto; + margin-right: auto; +} + +hr.doublespaceincode { + height: 1pt; + visibility: hidden; + font-size: 10px; +} + +/* +code.br { + height: 5em; +} +*/ + +#main .citation { + color: rgb(70%,0%,0%); + text-decoration: underline; +} + +table.infrule { + border: 0px; + margin-left: 50px; + margin-top: .5em; + margin-bottom: 1.2em; +} + +td.infrule { + font-family: monospace; + text-align: center; + /* color: rgb(35%,35%,70%); */ + padding: 0px; + line-height: 100%; +} + +tr.infrulemiddle hr { + margin: 1px 0 1px 0; +} + +.infrulenamecol { + color: rgb(60%,60%,60%); + font-size: 80%; + padding-left: 1em; + padding-bottom: 0.1em +} + +#footer { + font-size: 65%; + font-family: sans-serif; +} + +.id { display: inline; } + +.id[title="constructor"] { + color: #697f2f; +} + +.id[title="var"], +.id[title="variable"] { + color: rgb(40%,0%,40%); +} + +.id[title="definition"] { + color: rgb(0%,40%,0%); +} + +.id[title="abbreviation"] { + color: rgb(0%,40%,0%); +} + +.id[title="lemma"] { + color: rgb(0%,40%,0%); +} + +.id[title="instance"] { + color: rgb(0%,40%,0%); +} + +.id[title="projection"] { + color: rgb(0%,40%,0%); +} + +.id[title="method"] { + color: rgb(0%,40%,0%); +} + +.id[title="inductive"] { + color: #034764; +} + +.id[title="record"] { + color: rgb(0%,0%,80%); +} + +.id[title="class"] { + color: rgb(0%,0%,80%); +} + +.id[title="keyword"] { + color : #697f2f; + /* color: black; */ +} + +.inlinecode .id { + color: rgb(0%,0%,0%); +} + +.nowrap { + white-space: nowrap; +} + +.HIDEFROMHTML { + display: none; +} + +/* TOC */ + +#toc h2 { +/* padding-top: 13px; */ + padding-bottom: 13px; + padding-left: 8px; + margin-top: 5px; + margin-top: 20px; + /* OLD: padding: 10px; + line-height: 120%; + background-color: rgb(60%,60%,100%); */ +} + +#toc h2.ui-accordion-header { + padding: .5em .5em .5em .7em; + outline: none; +} + +#toc .ui-accordion .ui-accordion-content { + padding: 0.5em 2.5em 0.8em .9em; + border-top: 0; + margin-bottom: 1em; + /* bottom rule */ + border: none; + border-bottom: 1px solid transparent; + transition: border-bottom-color 0.25s ease-in; + transition-delay: 0.15s; +} + +#toc .ui-accordion .ui-accordion-content-active { + border-bottom: 1px solid #9b9b9b; + transition-delay: 0s; +} + +#toc h2.ui-accordion-header-active { + background: silver !important; +} + +#toc h2:not(.ui-accordion-header-active):hover { + background: rgba(0,0,0,0.04) !important; +} + +#toc h2 a:hover { + text-decoration: underline; +} + +#toc h2:hover::after { + content: "expand ▾"; + font-size: 80%; + float: right; + margin-top: 0.2em; + color: silver; + opacity: 1; + transition: opacity .5s ease-in-out; +} + +#toc h2.ui-accordion-header-active:hover::after { + opacity: 0; +} + +#toc h2 .select { background-image: url('media/image/arrow_down.jpg'); } +div#sec1.hide { display: none; } + +#toc ul { + padding-top: 8px; + font-size: 14px; + padding-left: 0; +} + +#toc ul ul { + margin-bottom: -8px; +} + +#toc li { + font-weight: 600; + list-style-type: none; + padding-top: 12px; + padding-bottom: 8px; +} + +#toc li li { + font-weight: 300; + list-style-type: circle; + padding-bottom: 3px; + padding-top: 0px; + margin-left: 19px; +} + + + + +/* Accordion Overrides */ + +/* Widget Bar */ +.ui-state-default, +.ui-widget-content .ui-state-default, +.ui-widget-header .ui-state-default, +.ui-button, +/* We use html here because we need a greater specificity to make sure disabled + works properly when clicked or hovered */ +html .ui-button.ui-state-disabled:hover, +html .ui-button.ui-state-disabled:active { + border: none!important; + /* BCP 3/17: I like it better without the rules... + border-bottom: 1px solid silver!important; */ + background: white !important; + font-weight: normal; + color: #454545!important; + font-weight: 400!important; + margin-top: 0px!important; + +} + +/* Misc visuals +----------------------------------*/ + +/* Corner radius */ +.ui-corner-all, +.ui-corner-top, +.ui-corner-left, +.ui-corner-tl { + border-top-left-radius: 0px!important; +} + +.ui-corner-all, +.ui-corner-top, +.ui-corner-right, +.ui-corner-tr { + border-top-right-radius: 0px!important; +} + +.ui-corner-all, +.ui-corner-bottom, +.ui-corner-left, +.ui-corner-bl { + border-bottom-left-radius: 0px!important; +} + +.ui-corner-all, +.ui-corner-bottom, +.ui-corner-right, +.ui-corner-br { + border-bottom-right-radius: 0px!important; +} + +html .ui-button.ui-state-disabled:focus { + color: red!important; +} + +/* Remove Icon */ +.ui-icon { display: none!important; } + +/* Widget */ +.ui-widget-content { + border: 1px solid #9e9e9e; + border-bottom-color: #b2b2b2; +} + +.ui-widget-content { + background: #ffffff; + color: #333333; +} + + +/* Index */ + +#index { + margin: 0; + padding: 0; + width: 100%; + font-style : normal; +} + +#index #frontispiece { + margin: auto; + padding: 1em; + width: 700px; +} + +.booktitle { + font-size : 300%; line-height: 100%; font-style:bold; + color: white; + padding-top: 70px; + padding-bottom: 20px; } +.authors { font-size : 200%; + line-height: 115%; } +.moreauthors { font-size : 170% } +.buttons { font-size : 170%; + margin-left: auto; + margin-right: auto; + font-style : bold; + } + +/* Link colors never changes */ + +A:link, A:visited, A:active, A:hover { + text-decoration: none; + color: #555555 +} + +/* Special color for the chapter header */ + +.booktitleinheader A:visited, .booktitleinheader A:active, .booktitleinheader A:hover, .booktitleinheader A:link { + text-decoration: none; + color: black; +} + +#index #entrance { + text-align: center; +} + +/* This was removed via CSS but the HTML is still generated */ +#index #footer { + display: none; +} + +.paragraph { + height: 0.6em; +} + +ul.doclist { + margin-top: 0em; + margin-bottom: 0em; +} + +/* Index styles */ + +/* Styles the author box (Intro class) and With (column class) */ + +.column { + float:left; + width: 43%; + margin:0 10px; + text-align: left; + font-size: 15px; + line-height: 25px; + padding-right: 20px; + min-height: 340px; +} + +.smallauthors { + font-size: 19px; + line-height: 25px; +} + +.mediumauthors { + font-size: 23px; + line-height: 33px; +} + +.largeauthors { + font-size: 28px; + line-height: 40px; +} + +.intro { + width: 35%; + font-size: 21px; + line-height: 27px; + font-weight: 600; + padding-right: 20px; +} + +.column.pub { + width: 40%; + margin-bottom: 20px; +} + +#index_content { + width: 100%!important; + display: block; + min-height: 400px; +} + +div.column.pub table tbody tr td { + text-align: center; padding: 10px; +} +div.column.pub table tbody tr td p { + text-align: left; + margin-top: 0; + font-weight: 600; + font-size: 13px!important; + line-height: 18px; +} + +/* Tables */ + +td.tab { + height: 16px; + font-weight: 600; + padding-left: 5px; + text-align: left!important; +} + +/* Styles tables on the index -- body class sf_home is used there */ + +body.sf_home table { + min-height: 450px; + vertical-align: top; +} + +body.sf_home table td { + vertical-align: top; + +} +body.sf_home table td p { + min-height: 100px; + +} + +table.logical { background-color: rgba(144, 160, 209, 0.5); } +table.logical tbody tr td.tab { background-color: #91a1d1; } + +table.language_found { background-color: rgba(178, 88, 88, 0.5); } +table.language_found tbody tr td.tab { background-color: #b25959; } + +table.algo { background-color: rgba(194, 194, 108, 0.5); } +table.algo tbody tr td.tab { background-color: #c2c26c; } + +table.qc { background-color: rgba(185, 170, 185, 0.5); } +table.qc tbody tr td.tab { background-color: #8b7d95; } + +table.vc { background-color: rgba(159, 125, 140, 0.5); } +table.vc tbody tr td.tab { background-color: rgb(159, 125, 140); } + +table.slf { background-color: rgba(219, 178, 127, 0.5); } +table.slf tbody tr td.tab { background-color: rgb(219, 178, 127); } + +/* Suggested background color for next volume */ +/* #c07d62 */ + +.ui-draggable, .ui-droppable { + background-position: top; +} + +/* Chapter dependencies (SVG) */ +.deps a polygon:hover { opacity: 0.6; stroke-width: 2; } +.deps a text { pointer-events: none; } + +/* A few specific things for the top-level SF landing page */ + +body.sf_home {background-color: #ededed; background-image: url(../media/image/core_mem_bg.jpg); } + +body.sf_home #header { + background-image: url(../media/image/core_mem_hdr_bg.jpg); + padding-bottom: 20px; +} + +body.sf_home #main_home { + background-color: transparent; +} + +/* A partial fix to a coqdoc bug... + See https://github.com/DeepSpec/sfdev/issues/236 */ +.inlinecode { white-space: pre; } +.inlinecode br { display: none; } + +#header { background-color: rgba(190, 170, 190, 0.5); } + +/* This volume's color */ +.section, ul#menu li.section_name, div.button { background-color: #8b7d95; } + +.slide img { + border: 2px solid gray; + margin: 1em; +} diff --git a/doc/res/fdl.org b/doc/res/fdl.org new file mode 100644 index 0000000..81e2cd9 --- /dev/null +++ b/doc/res/fdl.org @@ -0,0 +1,489 @@ +# The GNU Free Documentation License. +#+begin_center +Version 1.3, 3 November 2008 +#+end_center + +# This file is intended to be included within another document. + +#+begin_verse +Copyright \copy{} 2000, 2001, 2002, 2007, 2008 Free Software Foundation, Inc. +https://fsf.org/ + +Everyone is permitted to copy and distribute verbatim copies +of this license document, but changing it is not allowed. +#+end_verse + +0. [@0] PREAMBLE + + The purpose of this License is to make a manual, textbook, or other + functional and useful document @@texinfo:@dfn{@@free@@texinfo:}@@ + in the sense of freedom: to assure everyone the effective freedom + to copy and redistribute it, with or without modifying it, either + commercially or noncommercially. Secondarily, this License + preserves for the author and publisher a way to get credit for + their work, while not being considered responsible for + modifications made by others. + + This License is a kind of "copyleft", which means that derivative + works of the document must themselves be free in the same sense. + It complements the GNU General Public License, which is a copyleft + license designed for free software. + + We have designed this License in order to use it for manuals for + free software, because free software needs free documentation: + a free program should come with manuals providing the same freedoms + that the software does. But this License is not limited to + software manuals; it can be used for any textual work, regardless + of subject matter or whether it is published as a printed book. We + recommend this License principally for works whose purpose is + instruction or reference. + +1. APPLICABILITY AND DEFINITIONS + + This License applies to any manual or other work, in any medium, + that contains a notice placed by the copyright holder saying it can + be distributed under the terms of this License. Such a notice + grants a world-wide, royalty-free license, unlimited in duration, + to use that work under the conditions stated herein. The + "Document", below, refers to any such manual or work. Any member + of the public is a licensee, and is addressed as "you". You accept + the license if you copy, modify or distribute the work in a way + requiring permission under copyright law. + + A "Modified Version" of the Document means any work containing the + Document or a portion of it, either copied verbatim, or with + modifications and/or translated into another language. + + A "Secondary Section" is a named appendix or a front-matter section + of the Document that deals exclusively with the relationship of the + publishers or authors of the Document to the Document's overall + subject (or to related matters) and contains nothing that could + fall directly within that overall subject. (Thus, if the Document + is in part a textbook of mathematics, a Secondary Section may not + explain any mathematics.) The relationship could be a matter of + historical connection with the subject or with related matters, or + of legal, commercial, philosophical, ethical or political position + regarding them. + + The "Invariant Sections" are certain Secondary Sections whose + titles are designated, as being those of Invariant Sections, in the + notice that says that the Document is released under this License. + If a section does not fit the above definition of Secondary then it + is not allowed to be designated as Invariant. The Document may + contain zero Invariant Sections. If the Document does not identify + any Invariant Sections then there are none. + + The "Cover Texts" are certain short passages of text that are + listed, as Front-Cover Texts or Back-Cover Texts, in the notice + that says that the Document is released under this License. + A Front-Cover Text may be at most 5 words, and a Back-Cover Text + may be at most 25 words. + + A "Transparent" copy of the Document means a machine-readable copy, + represented in a format whose specification is available to the + general public, that is suitable for revising the document + straightforwardly with generic text editors or (for images composed + of pixels) generic paint programs or (for drawings) some widely + available drawing editor, and that is suitable for input to text + formatters or for automatic translation to a variety of formats + suitable for input to text formatters. A copy made in an otherwise + Transparent file format whose markup, or absence of markup, has + been arranged to thwart or discourage subsequent modification by + readers is not Transparent. An image format is not Transparent if + used for any substantial amount of text. A copy that is not + "Transparent" is called "Opaque". + + Examples of suitable formats for Transparent copies include plain + ASCII without markup, Texinfo input format, LaTeX input format, + SGML or XML using a publicly available DTD, and standard-conforming + simple HTML, PostScript or PDF designed for human modification. + Examples of transparent image formats include PNG, XCF and JPG. + Opaque formats include proprietary formats that can be read and + edited only by proprietary word processors, SGML or XML for which + the DTD and/or processing tools are not generally available, and + the machine-generated HTML, PostScript or PDF produced by some word + processors for output purposes only. + + The "Title Page" means, for a printed book, the title page itself, + plus such following pages as are needed to hold, legibly, the + material this License requires to appear in the title page. For + works in formats which do not have any title page as such, "Title + Page" means the text near the most prominent appearance of the + work's title, preceding the beginning of the body of the text. + + The "publisher" means any person or entity that distributes copies + of the Document to the public. + + A section "Entitled XYZ" means a named subunit of the Document + whose title either is precisely XYZ or contains XYZ in parentheses + following text that translates XYZ in another language. (Here XYZ + stands for a specific section name mentioned below, such as + "Acknowledgements", "Dedications", "Endorsements", or "History".) + To "Preserve the Title" of such a section when you modify the + Document means that it remains a section "Entitled XYZ" according + to this definition. + + The Document may include Warranty Disclaimers next to the notice + which states that this License applies to the Document. These + Warranty Disclaimers are considered to be included by reference in + this License, but only as regards disclaiming warranties: any other + implication that these Warranty Disclaimers may have is void and + has no effect on the meaning of this License. + +2. VERBATIM COPYING + + You may copy and distribute the Document in any medium, either + commercially or noncommercially, provided that this License, the + copyright notices, and the license notice saying this License + applies to the Document are reproduced in all copies, and that you + add no other conditions whatsoever to those of this License. You + may not use technical measures to obstruct or control the reading + or further copying of the copies you make or distribute. However, + you may accept compensation in exchange for copies. If you + distribute a large enough number of copies you must also follow the + conditions in section 3. + + You may also lend copies, under the same conditions stated above, + and you may publicly display copies. + +3. COPYING IN QUANTITY + + If you publish printed copies (or copies in media that commonly + have printed covers) of the Document, numbering more than 100, and + the Document's license notice requires Cover Texts, you must + enclose the copies in covers that carry, clearly and legibly, all + these Cover Texts: Front-Cover Texts on the front cover, and + Back-Cover Texts on the back cover. Both covers must also clearly + and legibly identify you as the publisher of these copies. The + front cover must present the full title with all words of the title + equally prominent and visible. You may add other material on the + covers in addition. Copying with changes limited to the covers, as + long as they preserve the title of the Document and satisfy these + conditions, can be treated as verbatim copying in other respects. + + If the required texts for either cover are too voluminous to fit + legibly, you should put the first ones listed (as many as fit + reasonably) on the actual cover, and continue the rest onto + adjacent pages. + + If you publish or distribute Opaque copies of the Document + numbering more than 100, you must either include a machine-readable + Transparent copy along with each Opaque copy, or state in or with + each Opaque copy a computer-network location from which the general + network-using public has access to download using public-standard + network protocols a complete Transparent copy of the Document, free + of added material. If you use the latter option, you must take + reasonably prudent steps, when you begin distribution of Opaque + copies in quantity, to ensure that this Transparent copy will + remain thus accessible at the stated location until at least one + year after the last time you distribute an Opaque copy (directly or + through your agents or retailers) of that edition to the public. + + It is requested, but not required, that you contact the authors of + the Document well before redistributing any large number of copies, + to give them a chance to provide you with an updated version of the + Document. + +4. MODIFICATIONS + + You may copy and distribute a Modified Version of the Document + under the conditions of sections 2 and 3 above, provided that you + release the Modified Version under precisely this License, with the + Modified Version filling the role of the Document, thus licensing + distribution and modification of the Modified Version to whoever + possesses a copy of it. In addition, you must do these things in + the Modified Version: + + #+attr_texinfo: :enum A + 1. Use in the Title Page (and on the covers, if any) a title + distinct from that of the Document, and from those of previous + versions (which should, if there were any, be listed in the + History section of the Document). You may use the same title as + a previous version if the original publisher of that version + gives permission. + + 2. List on the Title Page, as authors, one or more persons or + entities responsible for authorship of the modifications in the + Modified Version, together with at least five of the principal + authors of the Document (all of its principal authors, if it has + fewer than five), unless they release you from this requirement. + + 3. State on the Title page the name of the publisher of the + Modified Version, as the publisher. + + 4. Preserve all the copyright notices of the Document. + + 5. Add an appropriate copyright notice for your modifications + adjacent to the other copyright notices. + + 6. Include, immediately after the copyright notices, a license + notice giving the public permission to use the Modified Version + under the terms of this License, in the form shown in the + Addendum below. + + 7. Preserve in that license notice the full lists of Invariant + Sections and required Cover Texts given in the Document's + license notice. + + 8. Include an unaltered copy of this License. + + 9. Preserve the section Entitled "History", Preserve its Title, and + add to it an item stating at least the title, year, new authors, + and publisher of the Modified Version as given on the Title + Page. If there is no section Entitled "History" in the Document, + create one stating the title, year, authors, and publisher of + the Document as given on its Title Page, then add an item + describing the Modified Version as stated in the previous + sentence. + + 10. Preserve the network location, if any, given in the Document + for public access to a Transparent copy of the Document, and + likewise the network locations given in the Document for + previous versions it was based on. These may be placed in the + "History" section. You may omit a network location for a work + that was published at least four years before the Document + itself, or if the original publisher of the version it refers + to gives permission. + + 11. For any section Entitled "Acknowledgements" or "Dedications", + Preserve the Title of the section, and preserve in the section + all the substance and tone of each of the contributor + acknowledgements and/or dedications given therein. + + 12. Preserve all the Invariant Sections of the Document, unaltered + in their text and in their titles. Section numbers or the + equivalent are not considered part of the section titles. + + 13. Delete any section Entitled "Endorsements". Such a section may + not be included in the Modified Version. + + 14. Do not retitle any existing section to be Entitled + "Endorsements" or to conflict in title with any Invariant + Section. + + 15. Preserve any Warranty Disclaimers. + + If the Modified Version includes new front-matter sections or + appendices that qualify as Secondary Sections and contain no material + copied from the Document, you may at your option designate some or all + of these sections as invariant. To do this, add their titles to the + list of Invariant Sections in the Modified Version's license notice. + These titles must be distinct from any other section titles. + + You may add a section Entitled "Endorsements", provided it contains + nothing but endorsements of your Modified Version by various + parties---for example, statements of peer review or that the text has + been approved by an organization as the authoritative definition of a + standard. + + You may add a passage of up to five words as a Front-Cover Text, and a + passage of up to 25 words as a Back-Cover Text, to the end of the list + of Cover Texts in the Modified Version. Only one passage of + Front-Cover Text and one of Back-Cover Text may be added by (or + through arrangements made by) any one entity. If the Document already + includes a cover text for the same cover, previously added by you or + by arrangement made by the same entity you are acting on behalf of, + you may not add another; but you may replace the old one, on explicit + permission from the previous publisher that added the old one. + + The author(s) and publisher(s) of the Document do not by this License + give permission to use their names for publicity for or to assert or + imply endorsement of any Modified Version. + +5. COMBINING DOCUMENTS + + You may combine the Document with other documents released under + this License, under the terms defined in section 4 above for + modified versions, provided that you include in the combination all + of the Invariant Sections of all of the original documents, + unmodified, and list them all as Invariant Sections of your + combined work in its license notice, and that you preserve all + their Warranty Disclaimers. + + The combined work need only contain one copy of this License, and + multiple identical Invariant Sections may be replaced with a single + copy. If there are multiple Invariant Sections with the same name + but different contents, make the title of each such section unique + by adding at the end of it, in parentheses, the name of the + original author or publisher of that section if known, or else + a unique number. Make the same adjustment to the section titles in + the list of Invariant Sections in the license notice of the + combined work. + + In the combination, you must combine any sections Entitled + "History" in the various original documents, forming one section + Entitled "History"; likewise combine any sections Entitled + "Acknowledgements", and any sections Entitled "Dedications". You + must delete all sections Entitled "Endorsements." + +6. COLLECTIONS OF DOCUMENTS + + You may make a collection consisting of the Document and other + documents released under this License, and replace the individual + copies of this License in the various documents with a single copy + that is included in the collection, provided that you follow the + rules of this License for verbatim copying of each of the documents + in all other respects. + + You may extract a single document from such a collection, and + distribute it individually under this License, provided you insert + a copy of this License into the extracted document, and follow this + License in all other respects regarding verbatim copying of that + document. + +7. AGGREGATION WITH INDEPENDENT WORKS + + A compilation of the Document or its derivatives with other + separate and independent documents or works, in or on a volume of + a storage or distribution medium, is called an "aggregate" if the + copyright resulting from the compilation is not used to limit the + legal rights of the compilation's users beyond what the individual + works permit. When the Document is included in an aggregate, this + License does not apply to the other works in the aggregate which + are not themselves derivative works of the Document. + + If the Cover Text requirement of section 3 is applicable to these + copies of the Document, then if the Document is less than one half + of the entire aggregate, the Document's Cover Texts may be placed + on covers that bracket the Document within the aggregate, or the + electronic equivalent of covers if the Document is in electronic + form. Otherwise they must appear on printed covers that bracket + the whole aggregate. + +8. TRANSLATION + + Translation is considered a kind of modification, so you may + distribute translations of the Document under the terms of + section 4. Replacing Invariant Sections with translations requires + special permission from their copyright holders, but you may + include translations of some or all Invariant Sections in addition + to the original versions of these Invariant Sections. You may + include a translation of this License, and all the license notices + in the Document, and any Warranty Disclaimers, provided that you + also include the original English version of this License and the + original versions of those notices and disclaimers. In case of + a disagreement between the translation and the original version of + this License or a notice or disclaimer, the original version will + prevail. + + If a section in the Document is Entitled "Acknowledgements", + "Dedications", or "History", the requirement (section 4) to + Preserve its Title (section 1) will typically require changing the + actual title. + +9. TERMINATION + + You may not copy, modify, sublicense, or distribute the Document + except as expressly provided under this License. Any attempt + otherwise to copy, modify, sublicense, or distribute it is void, + and will automatically terminate your rights under this License. + + However, if you cease all violation of this License, then your + license from a particular copyright holder is reinstated (a) + provisionally, unless and until the copyright holder explicitly and + finally terminates your license, and (b) permanently, if the + copyright holder fails to notify you of the violation by some + reasonable means prior to 60 days after the cessation. + + Moreover, your license from a particular copyright holder is + reinstated permanently if the copyright holder notifies you of the + violation by some reasonable means, this is the first time you have + received notice of violation of this License (for any work) from + that copyright holder, and you cure the violation prior to 30 days + after your receipt of the notice. + + Termination of your rights under this section does not terminate + the licenses of parties who have received copies or rights from you + under this License. If your rights have been terminated and not + permanently reinstated, receipt of a copy of some or all of the + same material does not give you any rights to use it. + +10. FUTURE REVISIONS OF THIS LICENSE + + The Free Software Foundation may publish new, revised versions of + the GNU Free Documentation License from time to time. Such new + versions will be similar in spirit to the present version, but may + differ in detail to address new problems or concerns. See + https://www.gnu.org/copyleft/. + + Each version of the License is given a distinguishing version + number. If the Document specifies that a particular numbered + version of this License "or any later version" applies to it, you + have the option of following the terms and conditions either of + that specified version or of any later version that has been + published (not as a draft) by the Free Software Foundation. If + the Document does not specify a version number of this License, + you may choose any version ever published (not as a draft) by the + Free Software Foundation. If the Document specifies that a proxy + can decide which future versions of this License can be used, that + proxy's public statement of acceptance of a version permanently + authorizes you to choose that version for the Document. + +11. RELICENSING + + "Massive Multiauthor Collaboration Site" (or "MMC Site") means any + World Wide Web server that publishes copyrightable works and also + provides prominent facilities for anybody to edit those works. + A public wiki that anybody can edit is an example of such + a server. A "Massive Multiauthor Collaboration" (or "MMC") + contained in the site means any set of copyrightable works thus + published on the MMC site. + + "CC-BY-SA" means the Creative Commons Attribution-Share Alike 3.0 + license published by Creative Commons Corporation, + a not-for-profit corporation with a principal place of business in + San Francisco, California, as well as future copyleft versions of + that license published by that same organization. + + "Incorporate" means to publish or republish a Document, in whole + or in part, as part of another Document. + + An MMC is "eligible for relicensing" if it is licensed under this + License, and if all works that were first published under this + License somewhere other than this MMC, and subsequently + incorporated in whole or in part into the MMC, (1) had no cover + texts or invariant sections, and (2) were thus incorporated prior + to November 1, 2008. + + The operator of an MMC Site may republish an MMC contained in the + site under CC-BY-SA on the same site at any time before August 1, + 2009, provided the MMC is eligible for relicensing. + +#+texinfo: @page + +* ADDENDUM: How to use this License for your documents +:PROPERTIES: +:UNNUMBERED: notoc +:END: + +To use this License in a document you have written, include a copy of +the License in the document and put the following copyright and +license notices just after the title page: + +#+begin_example + Copyright (C) YEAR YOUR NAME. + Permission is granted to copy, distribute and/or modify this document + under the terms of the GNU Free Documentation License, Version 1.3 + or any later version published by the Free Software Foundation; + with no Invariant Sections, no Front-Cover Texts, and no Back-Cover + Texts. A copy of the license is included in the section entitled ``GNU + Free Documentation License''. +#+end_example + +If you have Invariant Sections, Front-Cover Texts and Back-Cover Texts, +replace the "with...Texts."\nbsp{}line with this: + +#+begin_example + with the Invariant Sections being LIST THEIR TITLES, with + the Front-Cover Texts being LIST, and with the Back-Cover Texts + being LIST. +#+end_example + +If you have Invariant Sections without Cover Texts, or some other +combination of the three, merge those two alternatives to suit the +situation. + +If your document contains nontrivial examples of program code, we +recommend releasing these examples in parallel under your choice of +free software license, such as the GNU General Public License, to +permit their use in free software. diff --git a/doc/res/install-deps.el b/doc/res/install-deps.el new file mode 100644 index 0000000..09cf4ae --- /dev/null +++ b/doc/res/install-deps.el @@ -0,0 +1,11 @@ +(require 'package) +(package-initialize) +(add-to-list 'package-archives '("nongnu" . "https://elpa.nongnu.org/nongnu/") t) +(add-to-list 'package-archives '("melpa" . "https://melpa.org/packages/") t) +(add-to-list 'package-archives '("gnu-devel" . "https://elpa.gnu.org/devel/") t) +(package-refresh-contents) + +(package-install 'org) +(package-install 'org-contrib) +(package-install 'org-transclusion) +(package-install 'htmlize) diff --git a/doc/res/publish.el b/doc/res/publish.el new file mode 100644 index 0000000..c083eb0 --- /dev/null +++ b/doc/res/publish.el @@ -0,0 +1,23 @@ +(require 'package) +(package-initialize) + +(require 'org) +(require 'org-transclusion) +(require 'ox) +(require 'ox-html) +(require 'htmlize) +(require 'ox-texinfo) +(require 'ox-man) + +(setq org-transclusion-exclude-elements nil + org-html-head-include-default-style nil + org-html-head-include-scripts nil + org-html-postamble-format '(("en" "")) + org-html-postamble t + org-html-divs '((preamble "header" "header") + (content "article" "content") + (postamble "footer" "postamble")) + org-html-doctype "html5" + org-html-htmlize-output-type 'css) + +(org-transclusion-add-all) diff --git a/doc/vericert.rst b/doc/vericert.rst new file mode 100644 index 0000000..3c8eaeb --- /dev/null +++ b/doc/vericert.rst @@ -0,0 +1,530 @@ +=============== +Vericert Manual +=============== + + +Introduction +------------ + +Vericert translates C code into a hardware description language called Verilog, which can then be +synthesised into hardware, to be placed onto a field-programmable gate array (FPGA) or +application-specific integrated circuit (ASIC). + +.. _fig:design: + +.. figure:: /_static/images/toolflow.png + + Current design of Vericert, where HTL is an intermediate language representing a finite state + machine with data-path (FSMD) and Verilog is the target language. + +The design shown in Figure `fig:design`_ shows how Vericert leverages an existing verified C +compiler called `CompCert `_ to perform this translation. + +.. index:: vericert + +.. _building: + +Building Vericert +----------------- + + +Testing +~~~~~~~ + +To test out ``vericert`` you can try the following examples which are in the test folder using the +following: + +.. code:: shell + + ./bin/vericert test/loop.c -o loop.v + ./bin/vericert test/conditional.c -o conditional.v + ./bin/vericert test/add.c -o add.v + +Or by running the test suite using the following command: + +.. code:: shell + + make test + +.. _using-vericert: + +.. index:: vericert + +Using Vericert +-------------- + +Vericert can be used to translate a subset of C into Verilog. As a simple example, consider the +following C file (``main.c``): + +.. code:: C + + void matrix_multiply(int first[2][2], int second[2][2], int multiply[2][2]) { + int sum = 0; + for (int c = 0; c < 2; c++) { + for (int d = 0; d < 2; d++) { + for (int k = 0; k < 2; k++) { + sum = sum + first[c][k]*second[k][d]; + } + multiply[c][d] = sum; + sum = 0; + } + } + } + + int main() { + int f[2][2] = {{1, 2}, {3, 4}}; + int s[2][2] = {{5, 6}, {7, 8}}; + int m[2][2] = {{0, 0}, {0, 0}}; + + matrix_multiply(f, s, m); + return m[1][1]; + } + +It can be compiled using the following command, assuming that vericert is somewhere on the path. + +.. code:: shell + + vericert main.c -o main.v + +The Verilog file contains a top-level test-bench, which can be given to any Verilog simulator to +simulate the hardware, which should give the same result as executing the C code. Using `Icarus +Verilog `_ as an example: + +.. code:: shell + + iverilog -o main_v main.v + +When executing, it should therefore print the following: + +.. code:: shell + + $ ./main_v + finished: 50 + +This gives the same result as executing the C in the following way: + +.. code:: shell + + $ gcc -o main_c main.c + $ ./main_c + $ echo $? + 50 + +Man pages +~~~~~~~~~ + +.. _unreleased-features: + +Unreleased Features +------------------- + +The following are unreleased features in Vericert that are currently being worked on and have not +been completely proven correct yet. Currently this includes features such as: + +- `scheduling`_, + +- `operation-chaining`_, + +- `if-conversion`_, and + +- `functions`_. + +This page gives some preliminary information on how the features are implemented and how the proofs +for the features are being done. Once these features are properly implemented, they will be added +to the proper documentation. + +.. _scheduling: + +Scheduling +~~~~~~~~~~ + +Scheduling is an optimisation which is used to run various instructions in parallel that are +independent to each other. + +.. _operation-chaining: + +Operation Chaining +~~~~~~~~~~~~~~~~~~ + +Operation chaining is an optimisation that can be added on to scheduling and allows for the +sequential execution of instructions in a clock cycle, while executing other instructions in +parallel in the same clock cycle. + +.. _if-conversion: + +If-conversion +~~~~~~~~~~~~~ + +If-conversion is an optimisation which can turn code with simple control flow into a single block +(called a hyper-block), using predicated instructions. + +.. _functions: + +Functions +~~~~~~~~~ + +Functions are currently only inlined in Vericert, however, we are working on a proper interface to +integrate function calls into the hardware. + +.. _coq-style-guide: + +Coq Style Guide +--------------- + +This style guide was taken from `Silveroak `_, it outlines +code style for Coq code in this repository. There are certainly other valid strategies and opinions +on Coq code style; this is laid out purely in the name of consistency. For a visual example of the +style, see the `example`_ at the bottom of this file. + +.. _code-organization: + +Code organization +~~~~~~~~~~~~~~~~~ + +.. _legal-banner: + +Legal banner +^^^^^^^^^^^^ + +- Files should begin with a copyright/license banner, as shown in the example above. + +.. _import-statements: + +Import statements +^^^^^^^^^^^^^^^^^ + +- ``Require Import`` statements should all go at the top of the file, followed by file-wide ``Import`` + statements. + + - =Import=s often contain notations or typeclass instances that might override notations or + instances from another library, so it’s nice to highlight them separately. + +- One ``Require Import`` statement per line; it’s easier to scan that way. + +- ``Require Import`` statements should use “fully-qualified” names (e.g. ``Require Import + Coq.ZArith.ZArith`` instead of ``Require Import ZArith``). + + - Use the ``Locate`` command to find the fully-qualified name! + +- ``Require Import``’s should go in the following order: + + 1. Standard library dependencies (start with ``Coq.``) + + 2. External dependencies (anything outside the current project) + + 3. Same-project dependencies + +- ``Require Import``’s with the same root library (the name before the first ``.``) should be + grouped together. Within each root-library group, they should be in alphabetical order (so + ``Coq.Lists.List`` before ``Coq.ZArith.ZArith``). + +.. _notations-and-scopes: + +Notations and scopes +^^^^^^^^^^^^^^^^^^^^ + +- Any file-wide ``Local Open Scope``’s should come immediately after the =Import=s (see example). + + - Always use ``Local Open Scope``; just ``Open Scope`` will sneakily open the scope for those who + import your file. + +- Put notations in their own separate modules or files, so that those who import your file can + choose whether or not they want the notations. + + - Conflicting notations can cause a lot of headache, so it comes in very handy to leave this + flexibility! + +.. _formatting: + +Formatting +~~~~~~~~~~ + +.. _line-length: + +Line length +^^^^^^^^^^^ + +- Maximum line length 80 characters. + + - Many Coq IDE setups divide the screen in half vertically and use only half to display source + code, so more than 80 characters can be genuinely hard to read on a laptop. + +.. _whitespace-and-indentation: + +Whitespace and indentation +^^^^^^^^^^^^^^^^^^^^^^^^^^ + +- No trailing whitespace. + +- Spaces, not tabs. + +- Files should end with a newline. + + - Many editors do this automatically on save. + +- Colons may be either “English-spaced”, with no space before the colon and one space after (``x: + nat``) or “French-spaced”, with one space before and after (``x : nat``). + +- Default indentation is 2 spaces. + + - Keeping this small prevents complex proofs from being indented ridiculously far, and matches IDE + defaults. + +- Use 2-space indents if inserting a line break immediately after: + + - ``Proof.`` + + - ``fun <...> =>`` + + - ``forall <...>,`` + + - ``exists <....>,`` + +- The style for indenting arguments in function application depends on where you make a line + break. If you make the line break immediately after the function name, use a 2-space + indent. However, if you make it after one or more arguments, align the next line with the first + argument: + + .. code:: coq + + (Z.pow + 1 2) + (Z.pow 1 2 3 + 4 5 6) + +- ``Inductive`` cases should not be indented. Example: + + .. code:: coq + + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo + . + +- ``match`` or ``lazymatch`` cases should line up with the “m” in ``match`` or “l” in ``lazymatch``, + as in the following examples: + + .. code:: coq + + match x with + | 3 => true + | _ => false + end. + + lazymatch x with + | 3 => idtac + | _ => fail "Not equal to 3:" x + end. + + repeat match goal with + | _ => progress subst + | _ => reflexivity + end. + + do 2 lazymatch goal with + | |- context [eq] => idtac + end. + +.. _definitions-and-fixpoints: + +Definitions and Fixpoints +~~~~~~~~~~~~~~~~~~~~~~~~~ + +- It’s okay to leave the return type of ``Definition``’s and ``Fixpoint``’s implicit + (e.g. ``Definition x := 5`` instead of ``Definition x : nat := 5``) when the type is very simple + or obvious (for instance, the definition is in a file which deals exclusively with operations on + ``Z``). + +.. _inductives: + +Inductives +~~~~~~~~~~ + +- The ``.`` ending an ``Inductive`` can be either on the same line as the last case or on its own + line immediately below. That is, both of the following are acceptable: + + .. code:: coq + + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo + . + Inductive Foo : Type := + | FooA : Foo + | FooB : Foo. + +.. _lemmatheorem-statements: + +Lemma/Theorem statements +~~~~~~~~~~~~~~~~~~~~~~~~ + +- Generally, use ``Theorem`` for the most important, top-level facts you prove and ``Lemma`` for + everything else. + +- Insert a line break after the colon in the lemma statement. + +- Insert a line break after the comma for ``forall`` or ``exist`` quantifiers. + +- Implication arrows (``->``) should share a line with the previous hypothesis, not the following + one. + +- There is no need to make a line break after every ``->``; short preconditions may share a line. + +.. _proofs-and-tactics: + +Proofs and tactics +~~~~~~~~~~~~~~~~~~ + +- Use the ``Proof`` command (lined up vertically with ``Lemma`` or ``Theorem`` it corresponds to) to + open a proof, and indent the first line after it 2 spaces. + +- Very small proofs (where ``Proof. Qed.`` is <= 80 characters) can go all in one line. + +- When ending a proof, align the ending statement (``Qed``, ``Admitted``, etc.) with ``Proof``. + +- Avoid referring to autogenerated names (e.g. ``H0``, ``n0``). It’s okay to let Coq generate these + names, but you should not explicitly refer to them in your proof. So ``intros; my_solver`` is + fine, but ``intros; apply H1; my_solver`` is not fine. + + - You can force a non-autogenerated name by either putting the variable before the colon in the + lemma statement (``Lemma foo x : ...`` instead of ``Lemma foo : forall x, ...``), or by passing + arguments to ``intros`` (e.g. ``intros ? x`` to name the second argument ``x``) + +- This way, the proof won’t break when new hypotheses are added or autogenerated variable names + change. + +- Use curly braces ``{}`` for subgoals, instead of bullets. + +- *Never write tactics with more than one subgoal focused.* This can make the proof very confusing + to step through! If you have more than one subgoal, use curly braces. + +- Consider adding a comment after the opening curly brace that explains what case you’re in (see + example). + + - This is not necessary for small subgoals but can help show the major lines of reasoning in large + proofs. + +- If invoking a tactic that is expected to return multiple subgoals, use ``[ | ... | ]`` before the + ``.`` to explicitly specify how many subgoals you expect. + + - Examples: ``split; [ | ].`` ``induction z; [ | | ].`` + + - This helps make code more maintainable, because it fails immediately if your tactic no longer + solves as many subgoals as expected (or unexpectedly solves more). + +- If invoking a string of tactics (composed by ``;``) that will break the goal into multiple + subgoals and then solve all but one, still use ``[ ]`` to enforce that all but one goal is solved. + + - Example: ``split; try lia; [ ]``. + +- Tactics that consist only of ``repeat``-ing a procedure (e.g. ``repeat match``, ``repeat first``) + should factor out a single step of that procedure a separate tactic called ``_step``, + because the single-step version is much easier to debug. For instance: + + .. code:: coq + + Ltac crush_step := + match goal with + | _ => progress subst + | _ => reflexivity + end. + Ltac crush := repeat crush_step. + +.. _naming: + +Naming +~~~~~~ + +- Helper proofs about standard library datatypes should go in a module that is named to match the + standard library module (see example). + + - This makes the helper proofs look like standard-library ones, which is helpful for categorizing + them if they’re genuinely at the standard-library level of abstraction. + +- Names of modules should start with capital letters. + +- Names of inductives and their constructors should start with capital letters. + +- Names of other definitions/lemmas should be snake case. + +.. _example: + +Example +~~~~~~~ + +A small standalone Coq file that exhibits many of the style points. + +.. code:: coq + + (* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2021 Name + * + * + *) + + Require Import Coq.Lists.List. + Require Import Coq.micromega.Lia. + Require Import Coq.ZArith.ZArith. + Import ListNotations. + Local Open Scope Z_scope. + + (* Helper proofs about standard library integers (Z) go within [Module Z] so + that they match standard-library Z lemmas when used. *) + Module Z. + Lemma pow_3_r x : x ^ 3 = x * x * x. + Proof. lia. Qed. (* very short proofs can go all on one line *) + + Lemma pow_4_r x : x ^ 4 = x * x * x * x. + Proof. + change 4 with (Z.succ (Z.succ (Z.succ (Z.succ 0)))). + repeat match goal with + | _ => rewrite Z.pow_1_r + | _ => rewrite Z.pow_succ_r by lia + | |- context [x * (?a * ?b)] => + replace (x * (a * b)) with (a * b * x) by lia + | _ => reflexivity + end. + Qed. + End Z. + (* Now we can access the lemmas above as Z.pow_3_r and Z.pow_4_r, as if they + were in the ZArith library! *) + + Definition bar (x y : Z) := x ^ (y + 1). + + (* example with a painfully manual proof to show case formatting *) + Lemma bar_upper_bound : + forall x y a, + 0 <= x <= a -> 0 <= y -> + 0 <= bar x y <= a ^ (y + 1). + Proof. + (* avoid referencing autogenerated names by explicitly naming variables *) + intros x y a Hx Hy. revert y Hy x a Hx. + (* explicitly indicate # subgoals with [ | ... | ] if > 1 *) + cbv [bar]; refine (natlike_ind _ _ _); [ | ]. + { (* y = 0 *) + intros; lia. } + { (* y = Z.succ _ *) + intros. + rewrite Z.add_succ_l, Z.pow_succ_r by lia. + split. + { (* 0 <= bar x y *) + apply Z.mul_nonneg_nonneg; [ lia | ]. + apply Z.pow_nonneg; lia. } + { (* bar x y < a ^ y *) + rewrite Z.pow_succ_r by lia. + apply Z.mul_le_mono_nonneg; try lia; + [ apply Z.pow_nonneg; lia | ]. + (* For more flexible proofs, use match statements to find hypotheses + rather than referring to them by autogenerated names like H0. In this + case, we'll take any hypothesis that applies to and then solves the + goal. *) + match goal with H : _ |- _ => apply H; solve [auto] end. } } + Qed. + + (* Put notations in a separate module or file so that importers can + decide whether or not to use them. *) + Module BarNotations. + Infix "#" := bar (at level 40) : Z_scope. + Notation "x '##'" := (bar x x) (at level 40) : Z_scope. + End BarNotations. diff --git a/docs/Makefile b/docs/Makefile deleted file mode 100644 index 9d4f361..0000000 --- a/docs/Makefile +++ /dev/null @@ -1,39 +0,0 @@ -all: manual src man-html - -install-deps: - emacs --batch --load ./res/install-deps.el - -%.man: %.org - emacs --batch --file $< --load ./res/publish.el --funcall org-man-export-to-man - -%.html: %.org - emacs --batch --file $< --load ./res/publish.el --funcall org-html-export-to-html - -manual: - mkdir -p manual - emacs --batch --file documentation.org --load ./res/publish.el --funcall org-texinfo-export-to-texinfo - makeinfo --html --number-sections --no-split \ - --css-ref "https://www.gnu.org/software/emacs/manual.css" \ - vericert.texi -o ./manual/index.html - cp -r images ./manual/. - -man-html: man.html - mkdir -p man - cp man.html ./man/vericert.1.html - -vericert.1: man.man - cp $< $@ - -src: - $(MAKE) -C .. doc - cp -r ../html src - -upload: - tar caf docs.tar.xz manual man src - rsync docs.tar.xz "zk@leika.ymhg.org:~" - -.PHONY: all upload manual man src install-deps man-html - -clean: - rm -rf manual man src - rm -f docs.tar.xz diff --git a/docs/README.org b/docs/README.org deleted file mode 100644 index 4113ed9..0000000 --- a/docs/README.org +++ /dev/null @@ -1,7 +0,0 @@ -#+title: Vericert Documentation - -The documentation for [[https://github.com/ymherklotz/vericert][Vericert]], which is written in the [[/documentation.org][documentation.org]] file. - -* How to develop on documentation - -The documentation uses =hugo= to generate the website, diff --git a/docs/common.org b/docs/common.org deleted file mode 100644 index aa27264..0000000 --- a/docs/common.org +++ /dev/null @@ -1,15 +0,0 @@ -#+author: Yann Herklotz -#+email: git@ymhg.org - -#+macro: version 1.2.2 -#+macro: modified 2022-02-24 - -#+macro: base_url https://vericert.ymhg.org - -#+macro: link src_emacs-lisp[:results raw]{(ymhg/link "$1" "$2")} - -#+macro: texinfo_head (eval (if (eq org-export-current-backend 'texinfo) "#+exclude_tags: noexport_texinfo" "#+exclude_tags: onlyexport_texinfo")) -#+macro: latex_head (eval (if (eq org-export-current-backend 'latex) "#+exclude_tags: noexport_latex" "#+exclude_tags: onlyexport_latex")) -#+macro: hugo_head (eval (if (eq org-export-current-backend 'hugo) "#+exclude_tags: noexport_hugo" "#+exclude_tags: onlyexport_hugo")) -{{{texinfo_head}}} -{{{latex_head}}} diff --git a/docs/documentation.org b/docs/documentation.org deleted file mode 100644 index 85bc8fd..0000000 --- a/docs/documentation.org +++ /dev/null @@ -1,562 +0,0 @@ -#+title: Vericert Manual -#+subtitle: Release {{{version}}} -#+author: Yann Herklotz -#+email: git@ymhg.org -#+language: en - -* Introduction -:PROPERTIES: -:EXPORT_FILE_NAME: _index -:EXPORT_HUGO_SECTION: docs -:CUSTOM_ID: docs -:END: - -Vericert translates C code into a hardware description language called Verilog, which can then be -synthesised into hardware, to be placed onto a field-programmable gate array (FPGA) or -application-specific integrated circuit (ASIC). - -#+attr_html: :width "600px" -#+caption: Current design of Vericert, where HTL is an intermediate language representing a finite state machine with data-path (FSMD) and Verilog is the target language. -#+name: fig:design -[[./images/toolflow.png]] - -The design shown in Figure [[fig:design]] shows how Vericert leverages an existing verified C compiler -called [[https://compcert.org/compcert-C.html][CompCert]] to perform this translation. - -#+texinfo: @insertcopying - -* COPYING -:PROPERTIES: -:COPYING: t -:END: - -Copyright (C) 2019-2022 Yann Herklotz. - -#+begin_quote -Permission is granted to copy, distribute and/or modify this document -under the terms of the GNU Free Documentation License, Version 1.3 -or any later version published by the Free Software Foundation; -with no Invariant Sections, no Front-Cover Texts, and no Back-Cover -Texts. A copy of the license is included in the section entitled ``GNU -Free Documentation License''. -#+end_quote - -* Building Vericert -:PROPERTIES: -:EXPORT_FILE_NAME: building -:EXPORT_HUGO_SECTION: docs -:CUSTOM_ID: building -:END: - -#+transclude: [[file:~/projects/vericert/README.org::#building][file:../README.org::#building]] :only-contents :exclude-elements "headline property-drawer" -#+transclude: [[file:~/projects/vericert/README.org::#downloading-compcert][file:../README.org::#downloading-compcert]] :level 2 -#+transclude: [[file:~/projects/vericert/README.org::#setting-up-nix][file:../README.org::#setting-up-nix]] :level 2 -#+transclude: [[file:~/projects/vericert/README.org::#makefile-build][file:../README.org::#makefile-build]] :level 2 - -** Testing - -To test out ~vericert~ you can try the following examples which are in the test folder using the -following: - -#+begin_src shell -./bin/vericert test/loop.c -o loop.v -./bin/vericert test/conditional.c -o conditional.v -./bin/vericert test/add.c -o add.v -#+end_src - -Or by running the test suite using the following command: - -#+begin_src shell -make test -#+end_src - -* Using Vericert -:PROPERTIES: -:CUSTOM_ID: using-vericert -:END: - -Vericert can be used to translate a subset of C into Verilog. As a simple example, consider the -following C file (~main.c~): - -#+begin_src C -void matrix_multiply(int first[2][2], int second[2][2], int multiply[2][2]) { - int sum = 0; - for (int c = 0; c < 2; c++) { - for (int d = 0; d < 2; d++) { - for (int k = 0; k < 2; k++) { - sum = sum + first[c][k]*second[k][d]; - } - multiply[c][d] = sum; - sum = 0; - } - } -} - -int main() { - int f[2][2] = {{1, 2}, {3, 4}}; - int s[2][2] = {{5, 6}, {7, 8}}; - int m[2][2] = {{0, 0}, {0, 0}}; - - matrix_multiply(f, s, m); - return m[1][1]; -} -#+end_src - -It can be compiled using the following command, assuming that vericert is somewhere on the path. - -#+begin_src shell -vericert main.c -o main.v -#+end_src - -The Verilog file contains a top-level test-bench, which can be given to any Verilog simulator to -simulate the hardware, which should give the same result as executing the C code. Using [[http://iverilog.icarus.com/][Icarus -Verilog]] as an example: - -#+begin_src shell -iverilog -o main_v main.v -#+end_src - -When executing, it should therefore print the following: - -#+begin_src shell -$ ./main_v -finished: 50 -#+end_src - -This gives the same result as executing the C in the following way: - -#+begin_src shell -$ gcc -o main_c main.c -$ ./main_c -$ echo $? -50 -#+end_src - -** Man pages - -#+transclude: [[file:man.org][file:man.org]] :exclude-elements "keyword" :level 3 - -* Unreleased Features -:PROPERTIES: -:EXPORT_FILE_NAME: unreleased -:EXPORT_HUGO_SECTION: docs -:CUSTOM_ID: unreleased-features -:END: - -The following are unreleased features in Vericert that are currently being worked on and have not -been completely proven correct yet. Currently this includes features such as: - -- [[#scheduling][scheduling]], -- [[#operation-chaining][operation chaining]], -- [[#if-conversion][if-conversion]], and -- [[#functions][functions]]. - -This page gives some preliminary information on how the features are implemented and how the proofs -for the features are being done. Once these features are properly implemented, they will be added -to the proper documentation. - -** Scheduling -:PROPERTIES: -:CUSTOM_ID: scheduling -:END: -#+cindex: scheduling - -Scheduling is an optimisation which is used to run various instructions in parallel that are -independent to each other. - -** Operation Chaining -:PROPERTIES: -:CUSTOM_ID: operation-chaining -:END: - -Operation chaining is an optimisation that can be added on to scheduling and allows for the -sequential execution of instructions in a clock cycle, while executing other instructions in -parallel in the same clock cycle. - -** If-conversion -:PROPERTIES: -:CUSTOM_ID: if-conversion -:END: - -If-conversion is an optimisation which can turn code with simple control flow into a single block -(called a hyper-block), using predicated instructions. - -** Functions -:PROPERTIES: -:CUSTOM_ID: functions -:END: - -Functions are currently only inlined in Vericert, however, we are working on a proper interface to -integrate function calls into the hardware. - -* Coq Style Guide - :PROPERTIES: - :CUSTOM_ID: coq-style-guide - :EXPORT_FILE_NAME: coq-style-guide - :EXPORT_HUGO_SECTION: docs - :END: - -This style guide was taken from [[https://github.com/project-oak/silveroak][Silveroak]], it outlines code style for Coq code in this -repository. There are certainly other valid strategies and opinions on Coq code style; this is laid -out purely in the name of consistency. For a visual example of the style, see the [[#example][example]] at the -bottom of this file. - -** Code organization - :PROPERTIES: - :CUSTOM_ID: code-organization - :END: -*** Legal banner - :PROPERTIES: - :CUSTOM_ID: legal-banner - :END: - -- Files should begin with a copyright/license banner, as shown in the example above. - -*** Import statements - :PROPERTIES: - :CUSTOM_ID: import-statements - :END: - -- =Require Import= statements should all go at the top of the file, followed by file-wide =Import= - statements. - - - =Import=s often contain notations or typeclass instances that might override notations or - instances from another library, so it's nice to highlight them separately. - -- One =Require Import= statement per line; it's easier to scan that way. -- =Require Import= statements should use "fully-qualified" names (e.g. =Require Import - Coq.ZArith.ZArith= instead of =Require Import ZArith=). - - - Use the =Locate= command to find the fully-qualified name! - -- =Require Import='s should go in the following order: - - 1. Standard library dependencies (start with =Coq.=) - 2. External dependencies (anything outside the current project) - 3. Same-project dependencies - -- =Require Import='s with the same root library (the name before the first =.=) should be grouped - together. Within each root-library group, they should be in alphabetical order (so =Coq.Lists.List= - before =Coq.ZArith.ZArith=). - -*** Notations and scopes - :PROPERTIES: - :CUSTOM_ID: notations-and-scopes - :END: - -- Any file-wide =Local Open Scope='s should come immediately after the =Import=s (see example). - - - Always use =Local Open Scope=; just =Open Scope= will sneakily open the scope for those who import - your file. - -- Put notations in their own separate modules or files, so that those who import your file can - choose whether or not they want the notations. - - - Conflicting notations can cause a lot of headache, so it comes in very handy to leave this - flexibility! - -** Formatting - :PROPERTIES: - :CUSTOM_ID: formatting - :END: -*** Line length - :PROPERTIES: - :CUSTOM_ID: line-length - :END: - -- Maximum line length 80 characters. - - - Many Coq IDE setups divide the screen in half vertically and use only half to display source - code, so more than 80 characters can be genuinely hard to read on a laptop. - -*** Whitespace and indentation - :PROPERTIES: - :CUSTOM_ID: whitespace-and-indentation - :END: - -- No trailing whitespace. - -- Spaces, not tabs. - -- Files should end with a newline. - - - Many editors do this automatically on save. - -- Colons may be either "English-spaced", with no space before the colon and one space after (=x: nat=) - or "French-spaced", with one space before and after (=x : nat=). - -- Default indentation is 2 spaces. - - - Keeping this small prevents complex proofs from being indented ridiculously far, and matches IDE - defaults. - -- Use 2-space indents if inserting a line break immediately after: - - - =Proof.= - - =fun <...> =>= - - =forall <...>,= - - =exists <....>,= - -- The style for indenting arguments in function application depends on where you make a line - break. If you make the line break immediately after the function name, use a 2-space - indent. However, if you make it after one or more arguments, align the next line with the first - argument: - - #+begin_src coq - (Z.pow - 1 2) - (Z.pow 1 2 3 - 4 5 6) - #+end_src - -- =Inductive= cases should not be indented. Example: - - #+begin_src coq - Inductive Foo : Type := - | FooA : Foo - | FooB : Foo - . - #+end_src - -- =match= or =lazymatch= cases should line up with the "m" in =match= or "l" in =lazymatch=, as in the - following examples: - - #+begin_src coq - match x with - | 3 => true - | _ => false - end. - - lazymatch x with - | 3 => idtac - | _ => fail "Not equal to 3:" x - end. - - repeat match goal with - | _ => progress subst - | _ => reflexivity - end. - - do 2 lazymatch goal with - | |- context [eq] => idtac - end. - #+end_src - -** Definitions and Fixpoints - :PROPERTIES: - :CUSTOM_ID: definitions-and-fixpoints - :END: - -- It's okay to leave the return type of =Definition='s and =Fixpoint='s implicit (e.g. ~Definition x := 5~ - instead of ~Definition x : nat := 5~) when the type is very simple or obvious (for instance, the - definition is in a file which deals exclusively with operations on =Z=). - -** Inductives - :PROPERTIES: - :CUSTOM_ID: inductives - :END: - -- The =.= ending an =Inductive= can be either on the same line as the last case or on its own line - immediately below. That is, both of the following are acceptable: - - #+begin_src coq - Inductive Foo : Type := - | FooA : Foo - | FooB : Foo - . - Inductive Foo : Type := - | FooA : Foo - | FooB : Foo. - #+end_src - -** Lemma/Theorem statements - :PROPERTIES: - :CUSTOM_ID: lemmatheorem-statements - :END: - -- Generally, use =Theorem= for the most important, top-level facts you prove and =Lemma= for everything - else. -- Insert a line break after the colon in the lemma statement. -- Insert a line break after the comma for =forall= or =exist= quantifiers. -- Implication arrows (=->=) should share a line with the previous hypothesis, not the following one. -- There is no need to make a line break after every =->=; short preconditions may share a line. - -** Proofs and tactics - :PROPERTIES: - :CUSTOM_ID: proofs-and-tactics - :END: - -- Use the =Proof= command (lined up vertically with =Lemma= or =Theorem= it corresponds to) to open a - proof, and indent the first line after it 2 spaces. - -- Very small proofs (where =Proof. Qed.= is <= 80 characters) can go all in one line. - -- When ending a proof, align the ending statement (=Qed=, =Admitted=, etc.) with =Proof=. - -- Avoid referring to autogenerated names (e.g. =H0=, =n0=). It's okay to let Coq generate these names, - but you should not explicitly refer to them in your proof. So =intros; my_solver= is fine, but - =intros; apply H1; my_solver= is not fine. - - - You can force a non-autogenerated name by either putting the variable before the colon in the - lemma statement (=Lemma foo x : ...= instead of =Lemma foo : forall x, ...=), or by passing - arguments to =intros= (e.g. =intros ? x= to name the second argument =x=) - -- This way, the proof won't break when new hypotheses are added or autogenerated variable names - change. - -- Use curly braces ={}= for subgoals, instead of bullets. - -- /Never write tactics with more than one subgoal focused./ This can make the proof very confusing to - step through! If you have more than one subgoal, use curly braces. - -- Consider adding a comment after the opening curly brace that explains what case you're in (see - example). - - - This is not necessary for small subgoals but can help show the major lines of reasoning in large - proofs. - -- If invoking a tactic that is expected to return multiple subgoals, use =[ | ... | ]= before the =.= to - explicitly specify how many subgoals you expect. - - - Examples: =split; [ | ].= =induction z; [ | | ].= - - This helps make code more maintainable, because it fails immediately if your tactic no longer - solves as many subgoals as expected (or unexpectedly solves more). - -- If invoking a string of tactics (composed by =;=) that will break the goal into multiple subgoals - and then solve all but one, still use =[ ]= to enforce that all but one goal is solved. - - - Example: =split; try lia; [ ]=. - -- Tactics that consist only of =repeat=-ing a procedure (e.g. =repeat match=, =repeat first=) should - factor out a single step of that procedure a separate tactic called =_step=, because - the single-step version is much easier to debug. For instance: - - #+begin_src coq - Ltac crush_step := - match goal with - | _ => progress subst - | _ => reflexivity - end. - Ltac crush := repeat crush_step. - #+end_src - -** Naming - :PROPERTIES: - :CUSTOM_ID: naming - :END: - -- Helper proofs about standard library datatypes should go in a module that is named to match the - standard library module (see example). - - - This makes the helper proofs look like standard-library ones, which is helpful for categorizing - them if they're genuinely at the standard-library level of abstraction. - -- Names of modules should start with capital letters. -- Names of inductives and their constructors should start with capital letters. -- Names of other definitions/lemmas should be snake case. - -** Example - :PROPERTIES: - :CUSTOM_ID: example - :END: -A small standalone Coq file that exhibits many of the style points. - -#+begin_src coq -(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 2021 Name - * - * - *) - - Require Import Coq.Lists.List. - Require Import Coq.micromega.Lia. - Require Import Coq.ZArith.ZArith. - Import ListNotations. - Local Open Scope Z_scope. - - (* Helper proofs about standard library integers (Z) go within [Module Z] so - that they match standard-library Z lemmas when used. *) - Module Z. - Lemma pow_3_r x : x ^ 3 = x * x * x. - Proof. lia. Qed. (* very short proofs can go all on one line *) - - Lemma pow_4_r x : x ^ 4 = x * x * x * x. - Proof. - change 4 with (Z.succ (Z.succ (Z.succ (Z.succ 0)))). - repeat match goal with - | _ => rewrite Z.pow_1_r - | _ => rewrite Z.pow_succ_r by lia - | |- context [x * (?a * ?b)] => - replace (x * (a * b)) with (a * b * x) by lia - | _ => reflexivity - end. - Qed. - End Z. - (* Now we can access the lemmas above as Z.pow_3_r and Z.pow_4_r, as if they - were in the ZArith library! *) - - Definition bar (x y : Z) := x ^ (y + 1). - - (* example with a painfully manual proof to show case formatting *) - Lemma bar_upper_bound : - forall x y a, - 0 <= x <= a -> 0 <= y -> - 0 <= bar x y <= a ^ (y + 1). - Proof. - (* avoid referencing autogenerated names by explicitly naming variables *) - intros x y a Hx Hy. revert y Hy x a Hx. - (* explicitly indicate # subgoals with [ | ... | ] if > 1 *) - cbv [bar]; refine (natlike_ind _ _ _); [ | ]. - { (* y = 0 *) - intros; lia. } - { (* y = Z.succ _ *) - intros. - rewrite Z.add_succ_l, Z.pow_succ_r by lia. - split. - { (* 0 <= bar x y *) - apply Z.mul_nonneg_nonneg; [ lia | ]. - apply Z.pow_nonneg; lia. } - { (* bar x y < a ^ y *) - rewrite Z.pow_succ_r by lia. - apply Z.mul_le_mono_nonneg; try lia; - [ apply Z.pow_nonneg; lia | ]. - (* For more flexible proofs, use match statements to find hypotheses - rather than referring to them by autogenerated names like H0. In this - case, we'll take any hypothesis that applies to and then solves the - goal. *) - match goal with H : _ |- _ => apply H; solve [auto] end. } } - Qed. - - (* Put notations in a separate module or file so that importers can - decide whether or not to use them. *) - Module BarNotations. - Infix "#" := bar (at level 40) : Z_scope. - Notation "x '##'" := (bar x x) (at level 40) : Z_scope. - End BarNotations. -#+end_src - -* Index - Features -:PROPERTIES: -:CUSTOM_ID: cindex -:APPENDIX: t -:INDEX: cp -:DESCRIPTION: Key concepts & features -:END: - -* Export Setup :noexport: - -#+setupfile: common.org - -#+export_file_name: vericert.texi - -#+texinfo_dir_category: High-level synthesis tool -#+texinfo_dir_title: Vericert -#+texinfo_dir_desc: Formally verified high-level synthesis tool - -* GNU Free Documentation License -:PROPERTIES: -:appendix: t -:END: - -#+include: res/fdl.org diff --git a/docs/images/toolflow.pdf b/docs/images/toolflow.pdf deleted file mode 100644 index 5fcee67..0000000 Binary files a/docs/images/toolflow.pdf and /dev/null differ diff --git a/docs/images/toolflow.png b/docs/images/toolflow.png deleted file mode 100644 index 601d6be..0000000 Binary files a/docs/images/toolflow.png and /dev/null differ diff --git a/docs/images/toolflow.svg b/docs/images/toolflow.svg deleted file mode 100644 index 0d8f39f..0000000 --- a/docs/images/toolflow.svg +++ /dev/null @@ -1,1250 +0,0 @@ - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - - diff --git a/docs/man.org b/docs/man.org deleted file mode 100644 index cb6143f..0000000 --- a/docs/man.org +++ /dev/null @@ -1,89 +0,0 @@ -#+title: vericert -#+man_class_options: :section-id "1" -#+options: toc:nil num:nil -#+html_head_extra: - -* NAME - -vericert - A formally verified high-level synthesis tool. - -* SYNOPSYS - -*vericert* [ *OPTION* ]... [ *FILE* ]... - -* DESCRIPTION - -** HLS Options: - -- --no-hls :: Do not use HLS and generate standard flow -- --simulate :: Simulate the result with the Verilog semantics -- --debug-hls :: Add debug logic to the Verilog -- --initialise-stack :: initialise the stack to all 0s - -** HLS Optimisations: - -- -fschedule :: Schedule the resulting hardware [off] -- -fif-conversion :: If-conversion optimisation [off] - -** General options: - -- -stdlib :: Set the path of the Compcert run-time library -- -v :: Print external commands before invoking them -- -timings :: Show the time spent in various compiler passes -- -version :: Print the version string and exit -- -target :: Generate code for the given target -- -conf :: Read configuration from file -- @ :: Read command line options from - -** Tracing Options: - -- -dprepro :: Save C file after preprocessing in .i -- -dparse :: Save C file after parsing and elaboration in .parsed.c -- -dc :: Save generated C in .compcert.c -- -dclight :: Save generated Clight in .light.c -- -dcminor :: Save generated Cminor in .cm -- -drtl :: Save RTL at various optimization points in .rtl. -- -drtlblock :: Save RTLBlock .rtlblock -- -dhtl :: Save HTL before Verilog generation .htl -- -dltl :: Save LTL after register allocation in .ltl -- -dmach :: Save generated Mach code in .mach -- -dasm :: Save generated assembly in .s -- -dall :: Save all generated intermediate files in . -- -sdump :: Save info for post-linking validation in .json -- -o :: Generate output in - -** Diagnostic options: - -- -Wall :: Enable all warnings -- -W :: Enable the specific -- -Wno- :: Disable the specific -- -Werror :: Make all warnings into errors -- -Werror= :: Turn into an error -- -Wno-error= :: Turn into a warning even if -Werror is specified -- -Wfatal-errors :: Turn all errors into fatal errors aborting the compilation -- -fdiagnostics-color :: Turn on colored diagnostics -- -fno-diagnostics-color :: Turn of colored diagnostics -- -fmax-errors= :: Maximum number of errors to report -- -fdiagnostics-show-option :: Print the option name with mappable diagnostics -- -fno-diagnostics-show-option :: Turn of printing of options with mappable diagnostics - -* AUTHOR - -Written by Yann Herklotz, Michalis Pardalos, James Pollard, Nadesh Ramanathan and John Wickerson. - -* COPYRIGHT - -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 . diff --git a/docs/res/coqdoc.css b/docs/res/coqdoc.css deleted file mode 100644 index 5572aaa..0000000 --- a/docs/res/coqdoc.css +++ /dev/null @@ -1,867 +0,0 @@ -body { - padding: 0px 0px; - margin: 0px 0px; - padding-left: 1em; - background-color: white; - font-family: sans-serif; - background-repeat: no-repeat; - background-size: 100%; -} - -#page { - display: block; - padding: 0px; - margin: 0px; -} - -#header { - min-height: 100px; - max-width: 760px; - margin: 0 auto; - padding-left: 80px; - padding-right: 80px; - padding-top: 30px; -} - -#header h1 { - padding: 0; - margin: 0; -} - -/* Menu */ -ul#menu { - padding: 0; - display: block; - margin: auto; -} - -ul#menu li, div.button { - display: inline-block; - background-color: white; - padding: 5px; - font-size: .70em; - text-transform: uppercase; - width: 30%; - text-align: center; - font-weight: 600; -} - -div.button { - margin-top:5px; - width: 40%; -} - -#button_block {margin-top:50px;} - -ul#menu a.hover li { - background-color: red; -} - -/* Contents */ - -#main, #main_home { - display: block; - padding: 80px; - padding-top: 60px; /* BCP */ - font-size: 100%; - line-height: 100%; - max-width: 760px; - background-color: #ffffff; - margin: 0 auto; -} - -#main_home { - background-color: rgba(255, 255, 255, 0.5); -} - -#index_content div.intro p { - font-size: 12px; -} - -#main h1 { - /* line-height: 80%; */ - line-height: normal; - padding-top: 3px; - padding-bottom: 4px; - /* Demitri: font-size: 22pt; */ - font-size: 200%; /* BCP */ -} - -/* allow for multi-line headers */ -#main a.idref:visited {color : #416DFF; text-decoration : none; } -#main a.idref:link {color : #416DFF; text-decoration : none; } -#main a.idref:hover {text-decoration : none; } -#main a.idref:active {text-decoration : none; } - -#main a.modref:visited {color : #416DFF; text-decoration : none; } -#main a.modref:link {color : #416DFF; text-decoration : none; } -#main a.modref:hover {text-decoration : none; } -#main a.modref:active {text-decoration : none; } - -#main .keyword { color : #697f2f } -#main { color: black } - -/* General section class - applies to all .section IDs */ -.section { - padding-top: 12px; - padding-bottom: 11px; - padding-left: 8px; - margin-top: 5px; - margin-bottom: 3px; - margin-top: 18px; - font-size : 125%; - color: #ffffff; -} - -/* Book title in header */ -.booktitleinheader { - color: #000000; - text-transform: uppercase; - font-weight: 300; - letter-spacing: 1px; - font-size: 125%; - margin-left: 0px; - margin-bottom: 22px; - } - -/* Chapter titles */ -.libtitle { - max-width: 860px; - text-transform: uppercase; - margin: 5px auto; - font-weight: 500; - padding-bottom: 2px; - font-size: 120%; - letter-spacing: 3px; - } - -.subtitle { - display: block; - padding-top: 10px; - font-size: 70%; - line-height: normal; -} - -h2.section { - color: #2a2c71; - background-color: transparent; - padding-left: 0px; - padding-top: 0px; - padding-bottom: 0px; - /* margin-top: 0px; */ - margin-top: 9px; /* BCP 2/20 */ - font-size : 135%; } - -h3.section { - /* background-color: rgb(90%,90%,100%); */ - background-color: white; - /* padding-left: 8px; */ - padding-left: 0px; - /* padding-top: 7px; */ - padding-top: 0px; - /* padding-bottom: 0px; */ - padding-bottom: 0.5em; - /* margin-top: 9px; */ - /* margin-top: 4px; (BCP 2/20) */ - margin-top: 9px; /* BCP 2/20 */ - font-size : 115%; - color:black; -} - -h4.section { - margin-top: .2em; - background-color: white; - color: #2a2c71; - padding-left: 0px; - padding-top: 0px; - padding-bottom: 0.5em; /* 0px; */ - font-size : 100%; - font-style : bold; - text-decoration : underline; -} - -#header p { - font-size: 13px; -} - -/* Sets up Main ID and margins */ - -#main .doc { - margin: 0px auto; - font-size: 14px; - line-height: 22px; - /* max-width: 570px; */ - color: black; - /* text-align: justify; */ - border-style: plain; - /* This might work better after changing back to standard coqdoc: */ - padding-top: 10px; - /* padding-top: 18px; */ -} - -.quote { - margin-left: auto; - margin-right: auto; - width: 40em; - color: darkred; -} - -.loud { - color: darkred; -} - -pre { - margin-top: 0px; - margin-bottom: 0px; -} - -.inlinecode { - display: inline; - /* font-size: 125%; */ - color: #444444; - font-family: monospace } - -.doc .inlinecode { - display: inline; - font-size: 105%; - color: rgb(35%,35%,70%); - font-family: monospace } - -.doc .inlinecode .id { -/* I am changing this to white in style below: - color: rgb(35%,35%,70%); -*/ -} - -h1 .inlinecode .id, h1.section span.inlinecode { - color: #ffffff; -} - -.inlinecodenm { - display: inline; - /* font-size: 125%; */ - color: #444444; -} - -.doc .inlinecodenm { - display: inline; - color: rgb(35%,35%,70%); -} - -.doc .inlinecodenm .id { - color: rgb(35%,35%,70%); -} - -.doc .code { - display: inline; - font-size: 110%; - color: rgb(35%,35%,70%); - font-family: monospace; - padding-left: 0px; -} - -.comment { - display: inline; - font-family: monospace; - color: rgb(50%,50%,80%); -} - -.inlineref { - display: inline; - /* font-family: monospace; */ - color: rgb(50%,50%,80%); -} - -.show::before { - /* content: "more"; */ - content: "+"; -} - -.show { - background-color: rgb(93%,93%,93%); - display: inline; - font-family: monospace; - font-size: 60%; - padding-top: 1px; - padding-bottom: 2px; - padding-left: 4px; - padding-right: 4px; - color: rgb(60%,60%,60%); -} - -/* -.show { - display: inline; - font-family: monospace; - font-size: 60%; - padding-top: 0px; - padding-bottom: 0px; - padding-left: 10px; - border: 1px; - border-style: solid; - color: rgb(75%,75%,85%); -} -*/ - -.proofbox { - font-size: 90%; - color: rgb(75%,75%,75%); -} - -#main .less-space { - margin-top: -12px; -} - -/* Inline quizzes */ -.quiz:before { - color: rgb(40%,40%,40%); - /* content: "- Quick Check -" ; */ - display: block; - text-align: center; - margin-bottom: 5px; -} -.quiz { - border: 4px; - border-color: rgb(80%,80%,80%); - /* - margin-left: 40px; - margin-right: 100px; - */ - padding: 5px; - padding-left: 8px; - padding-right: 8px; - margin-top: 10px; - margin-bottom: 15px; - border-style: solid; -} - -/* For textual ones... */ -.show-old { - display: inline; - font-family: monospace; - font-size: 80%; - padding-top: 0px; - padding-bottom: 0px; - padding-left: 3px; - padding-right: 3px; - border: 1px; - margin-top: 5px; /* doesn't work?! */ - border-style: solid; - color: rgb(75%,75%,85%); -} - -.largebr { - margin-top: 10px; -} - -.code { - padding-left: 20px; - font-size: 14px; - font-family: monospace; - line-height: 17px; - margin-top: 9px; -} - -/* Working: -.code { - display: block; - padding-left: 0px; - font-size: 110%; - font-family: monospace; - } -*/ - -.code-space { - max-width: 50em; - margin-top: 0em; - /* margin-bottom: -.5em; */ - margin-left: auto; - margin-right: auto; -} - -.code-tight { - max-width: 50em; - margin-top: .6em; - /* margin-bottom: -.7em; */ - margin-left: auto; - margin-right: auto; -} - -hr.doublespaceincode { - height: 1pt; - visibility: hidden; - font-size: 10px; -} - -/* -code.br { - height: 5em; -} -*/ - -#main .citation { - color: rgb(70%,0%,0%); - text-decoration: underline; -} - -table.infrule { - border: 0px; - margin-left: 50px; - margin-top: .5em; - margin-bottom: 1.2em; -} - -td.infrule { - font-family: monospace; - text-align: center; - /* color: rgb(35%,35%,70%); */ - padding: 0px; - line-height: 100%; -} - -tr.infrulemiddle hr { - margin: 1px 0 1px 0; -} - -.infrulenamecol { - color: rgb(60%,60%,60%); - font-size: 80%; - padding-left: 1em; - padding-bottom: 0.1em -} - -#footer { - font-size: 65%; - font-family: sans-serif; -} - -.id { display: inline; } - -.id[title="constructor"] { - color: #697f2f; -} - -.id[title="var"], -.id[title="variable"] { - color: rgb(40%,0%,40%); -} - -.id[title="definition"] { - color: rgb(0%,40%,0%); -} - -.id[title="abbreviation"] { - color: rgb(0%,40%,0%); -} - -.id[title="lemma"] { - color: rgb(0%,40%,0%); -} - -.id[title="instance"] { - color: rgb(0%,40%,0%); -} - -.id[title="projection"] { - color: rgb(0%,40%,0%); -} - -.id[title="method"] { - color: rgb(0%,40%,0%); -} - -.id[title="inductive"] { - color: #034764; -} - -.id[title="record"] { - color: rgb(0%,0%,80%); -} - -.id[title="class"] { - color: rgb(0%,0%,80%); -} - -.id[title="keyword"] { - color : #697f2f; - /* color: black; */ -} - -.inlinecode .id { - color: rgb(0%,0%,0%); -} - -.nowrap { - white-space: nowrap; -} - -.HIDEFROMHTML { - display: none; -} - -/* TOC */ - -#toc h2 { -/* padding-top: 13px; */ - padding-bottom: 13px; - padding-left: 8px; - margin-top: 5px; - margin-top: 20px; - /* OLD: padding: 10px; - line-height: 120%; - background-color: rgb(60%,60%,100%); */ -} - -#toc h2.ui-accordion-header { - padding: .5em .5em .5em .7em; - outline: none; -} - -#toc .ui-accordion .ui-accordion-content { - padding: 0.5em 2.5em 0.8em .9em; - border-top: 0; - margin-bottom: 1em; - /* bottom rule */ - border: none; - border-bottom: 1px solid transparent; - transition: border-bottom-color 0.25s ease-in; - transition-delay: 0.15s; -} - -#toc .ui-accordion .ui-accordion-content-active { - border-bottom: 1px solid #9b9b9b; - transition-delay: 0s; -} - -#toc h2.ui-accordion-header-active { - background: silver !important; -} - -#toc h2:not(.ui-accordion-header-active):hover { - background: rgba(0,0,0,0.04) !important; -} - -#toc h2 a:hover { - text-decoration: underline; -} - -#toc h2:hover::after { - content: "expand ▾"; - font-size: 80%; - float: right; - margin-top: 0.2em; - color: silver; - opacity: 1; - transition: opacity .5s ease-in-out; -} - -#toc h2.ui-accordion-header-active:hover::after { - opacity: 0; -} - -#toc h2 .select { background-image: url('media/image/arrow_down.jpg'); } -div#sec1.hide { display: none; } - -#toc ul { - padding-top: 8px; - font-size: 14px; - padding-left: 0; -} - -#toc ul ul { - margin-bottom: -8px; -} - -#toc li { - font-weight: 600; - list-style-type: none; - padding-top: 12px; - padding-bottom: 8px; -} - -#toc li li { - font-weight: 300; - list-style-type: circle; - padding-bottom: 3px; - padding-top: 0px; - margin-left: 19px; -} - - - - -/* Accordion Overrides */ - -/* Widget Bar */ -.ui-state-default, -.ui-widget-content .ui-state-default, -.ui-widget-header .ui-state-default, -.ui-button, -/* We use html here because we need a greater specificity to make sure disabled - works properly when clicked or hovered */ -html .ui-button.ui-state-disabled:hover, -html .ui-button.ui-state-disabled:active { - border: none!important; - /* BCP 3/17: I like it better without the rules... - border-bottom: 1px solid silver!important; */ - background: white !important; - font-weight: normal; - color: #454545!important; - font-weight: 400!important; - margin-top: 0px!important; - -} - -/* Misc visuals -----------------------------------*/ - -/* Corner radius */ -.ui-corner-all, -.ui-corner-top, -.ui-corner-left, -.ui-corner-tl { - border-top-left-radius: 0px!important; -} - -.ui-corner-all, -.ui-corner-top, -.ui-corner-right, -.ui-corner-tr { - border-top-right-radius: 0px!important; -} - -.ui-corner-all, -.ui-corner-bottom, -.ui-corner-left, -.ui-corner-bl { - border-bottom-left-radius: 0px!important; -} - -.ui-corner-all, -.ui-corner-bottom, -.ui-corner-right, -.ui-corner-br { - border-bottom-right-radius: 0px!important; -} - -html .ui-button.ui-state-disabled:focus { - color: red!important; -} - -/* Remove Icon */ -.ui-icon { display: none!important; } - -/* Widget */ -.ui-widget-content { - border: 1px solid #9e9e9e; - border-bottom-color: #b2b2b2; -} - -.ui-widget-content { - background: #ffffff; - color: #333333; -} - - -/* Index */ - -#index { - margin: 0; - padding: 0; - width: 100%; - font-style : normal; -} - -#index #frontispiece { - margin: auto; - padding: 1em; - width: 700px; -} - -.booktitle { - font-size : 300%; line-height: 100%; font-style:bold; - color: white; - padding-top: 70px; - padding-bottom: 20px; } -.authors { font-size : 200%; - line-height: 115%; } -.moreauthors { font-size : 170% } -.buttons { font-size : 170%; - margin-left: auto; - margin-right: auto; - font-style : bold; - } - -/* Link colors never changes */ - -A:link, A:visited, A:active, A:hover { - text-decoration: none; - color: #555555 -} - -/* Special color for the chapter header */ - -.booktitleinheader A:visited, .booktitleinheader A:active, .booktitleinheader A:hover, .booktitleinheader A:link { - text-decoration: none; - color: black; -} - -#index #entrance { - text-align: center; -} - -/* This was removed via CSS but the HTML is still generated */ -#index #footer { - display: none; -} - -.paragraph { - height: 0.6em; -} - -ul.doclist { - margin-top: 0em; - margin-bottom: 0em; -} - -/* Index styles */ - -/* Styles the author box (Intro class) and With (column class) */ - -.column { - float:left; - width: 43%; - margin:0 10px; - text-align: left; - font-size: 15px; - line-height: 25px; - padding-right: 20px; - min-height: 340px; -} - -.smallauthors { - font-size: 19px; - line-height: 25px; -} - -.mediumauthors { - font-size: 23px; - line-height: 33px; -} - -.largeauthors { - font-size: 28px; - line-height: 40px; -} - -.intro { - width: 35%; - font-size: 21px; - line-height: 27px; - font-weight: 600; - padding-right: 20px; -} - -.column.pub { - width: 40%; - margin-bottom: 20px; -} - -#index_content { - width: 100%!important; - display: block; - min-height: 400px; -} - -div.column.pub table tbody tr td { - text-align: center; padding: 10px; -} -div.column.pub table tbody tr td p { - text-align: left; - margin-top: 0; - font-weight: 600; - font-size: 13px!important; - line-height: 18px; -} - -/* Tables */ - -td.tab { - height: 16px; - font-weight: 600; - padding-left: 5px; - text-align: left!important; -} - -/* Styles tables on the index -- body class sf_home is used there */ - -body.sf_home table { - min-height: 450px; - vertical-align: top; -} - -body.sf_home table td { - vertical-align: top; - -} -body.sf_home table td p { - min-height: 100px; - -} - -table.logical { background-color: rgba(144, 160, 209, 0.5); } -table.logical tbody tr td.tab { background-color: #91a1d1; } - -table.language_found { background-color: rgba(178, 88, 88, 0.5); } -table.language_found tbody tr td.tab { background-color: #b25959; } - -table.algo { background-color: rgba(194, 194, 108, 0.5); } -table.algo tbody tr td.tab { background-color: #c2c26c; } - -table.qc { background-color: rgba(185, 170, 185, 0.5); } -table.qc tbody tr td.tab { background-color: #8b7d95; } - -table.vc { background-color: rgba(159, 125, 140, 0.5); } -table.vc tbody tr td.tab { background-color: rgb(159, 125, 140); } - -table.slf { background-color: rgba(219, 178, 127, 0.5); } -table.slf tbody tr td.tab { background-color: rgb(219, 178, 127); } - -/* Suggested background color for next volume */ -/* #c07d62 */ - -.ui-draggable, .ui-droppable { - background-position: top; -} - -/* Chapter dependencies (SVG) */ -.deps a polygon:hover { opacity: 0.6; stroke-width: 2; } -.deps a text { pointer-events: none; } - -/* A few specific things for the top-level SF landing page */ - -body.sf_home {background-color: #ededed; background-image: url(../media/image/core_mem_bg.jpg); } - -body.sf_home #header { - background-image: url(../media/image/core_mem_hdr_bg.jpg); - padding-bottom: 20px; -} - -body.sf_home #main_home { - background-color: transparent; -} - -/* A partial fix to a coqdoc bug... - See https://github.com/DeepSpec/sfdev/issues/236 */ -.inlinecode { white-space: pre; } -.inlinecode br { display: none; } - -#header { background-color: rgba(190, 170, 190, 0.5); } - -/* This volume's color */ -.section, ul#menu li.section_name, div.button { background-color: #8b7d95; } - -.slide img { - border: 2px solid gray; - margin: 1em; -} diff --git a/docs/res/fdl.org b/docs/res/fdl.org deleted file mode 100644 index 81e2cd9..0000000 --- a/docs/res/fdl.org +++ /dev/null @@ -1,489 +0,0 @@ -# The GNU Free Documentation License. -#+begin_center -Version 1.3, 3 November 2008 -#+end_center - -# This file is intended to be included within another document. - -#+begin_verse -Copyright \copy{} 2000, 2001, 2002, 2007, 2008 Free Software Foundation, Inc. -https://fsf.org/ - -Everyone is permitted to copy and distribute verbatim copies -of this license document, but changing it is not allowed. -#+end_verse - -0. [@0] PREAMBLE - - The purpose of this License is to make a manual, textbook, or other - functional and useful document @@texinfo:@dfn{@@free@@texinfo:}@@ - in the sense of freedom: to assure everyone the effective freedom - to copy and redistribute it, with or without modifying it, either - commercially or noncommercially. Secondarily, this License - preserves for the author and publisher a way to get credit for - their work, while not being considered responsible for - modifications made by others. - - This License is a kind of "copyleft", which means that derivative - works of the document must themselves be free in the same sense. - It complements the GNU General Public License, which is a copyleft - license designed for free software. - - We have designed this License in order to use it for manuals for - free software, because free software needs free documentation: - a free program should come with manuals providing the same freedoms - that the software does. But this License is not limited to - software manuals; it can be used for any textual work, regardless - of subject matter or whether it is published as a printed book. We - recommend this License principally for works whose purpose is - instruction or reference. - -1. APPLICABILITY AND DEFINITIONS - - This License applies to any manual or other work, in any medium, - that contains a notice placed by the copyright holder saying it can - be distributed under the terms of this License. Such a notice - grants a world-wide, royalty-free license, unlimited in duration, - to use that work under the conditions stated herein. The - "Document", below, refers to any such manual or work. Any member - of the public is a licensee, and is addressed as "you". You accept - the license if you copy, modify or distribute the work in a way - requiring permission under copyright law. - - A "Modified Version" of the Document means any work containing the - Document or a portion of it, either copied verbatim, or with - modifications and/or translated into another language. - - A "Secondary Section" is a named appendix or a front-matter section - of the Document that deals exclusively with the relationship of the - publishers or authors of the Document to the Document's overall - subject (or to related matters) and contains nothing that could - fall directly within that overall subject. (Thus, if the Document - is in part a textbook of mathematics, a Secondary Section may not - explain any mathematics.) The relationship could be a matter of - historical connection with the subject or with related matters, or - of legal, commercial, philosophical, ethical or political position - regarding them. - - The "Invariant Sections" are certain Secondary Sections whose - titles are designated, as being those of Invariant Sections, in the - notice that says that the Document is released under this License. - If a section does not fit the above definition of Secondary then it - is not allowed to be designated as Invariant. The Document may - contain zero Invariant Sections. If the Document does not identify - any Invariant Sections then there are none. - - The "Cover Texts" are certain short passages of text that are - listed, as Front-Cover Texts or Back-Cover Texts, in the notice - that says that the Document is released under this License. - A Front-Cover Text may be at most 5 words, and a Back-Cover Text - may be at most 25 words. - - A "Transparent" copy of the Document means a machine-readable copy, - represented in a format whose specification is available to the - general public, that is suitable for revising the document - straightforwardly with generic text editors or (for images composed - of pixels) generic paint programs or (for drawings) some widely - available drawing editor, and that is suitable for input to text - formatters or for automatic translation to a variety of formats - suitable for input to text formatters. A copy made in an otherwise - Transparent file format whose markup, or absence of markup, has - been arranged to thwart or discourage subsequent modification by - readers is not Transparent. An image format is not Transparent if - used for any substantial amount of text. A copy that is not - "Transparent" is called "Opaque". - - Examples of suitable formats for Transparent copies include plain - ASCII without markup, Texinfo input format, LaTeX input format, - SGML or XML using a publicly available DTD, and standard-conforming - simple HTML, PostScript or PDF designed for human modification. - Examples of transparent image formats include PNG, XCF and JPG. - Opaque formats include proprietary formats that can be read and - edited only by proprietary word processors, SGML or XML for which - the DTD and/or processing tools are not generally available, and - the machine-generated HTML, PostScript or PDF produced by some word - processors for output purposes only. - - The "Title Page" means, for a printed book, the title page itself, - plus such following pages as are needed to hold, legibly, the - material this License requires to appear in the title page. For - works in formats which do not have any title page as such, "Title - Page" means the text near the most prominent appearance of the - work's title, preceding the beginning of the body of the text. - - The "publisher" means any person or entity that distributes copies - of the Document to the public. - - A section "Entitled XYZ" means a named subunit of the Document - whose title either is precisely XYZ or contains XYZ in parentheses - following text that translates XYZ in another language. (Here XYZ - stands for a specific section name mentioned below, such as - "Acknowledgements", "Dedications", "Endorsements", or "History".) - To "Preserve the Title" of such a section when you modify the - Document means that it remains a section "Entitled XYZ" according - to this definition. - - The Document may include Warranty Disclaimers next to the notice - which states that this License applies to the Document. These - Warranty Disclaimers are considered to be included by reference in - this License, but only as regards disclaiming warranties: any other - implication that these Warranty Disclaimers may have is void and - has no effect on the meaning of this License. - -2. VERBATIM COPYING - - You may copy and distribute the Document in any medium, either - commercially or noncommercially, provided that this License, the - copyright notices, and the license notice saying this License - applies to the Document are reproduced in all copies, and that you - add no other conditions whatsoever to those of this License. You - may not use technical measures to obstruct or control the reading - or further copying of the copies you make or distribute. However, - you may accept compensation in exchange for copies. If you - distribute a large enough number of copies you must also follow the - conditions in section 3. - - You may also lend copies, under the same conditions stated above, - and you may publicly display copies. - -3. COPYING IN QUANTITY - - If you publish printed copies (or copies in media that commonly - have printed covers) of the Document, numbering more than 100, and - the Document's license notice requires Cover Texts, you must - enclose the copies in covers that carry, clearly and legibly, all - these Cover Texts: Front-Cover Texts on the front cover, and - Back-Cover Texts on the back cover. Both covers must also clearly - and legibly identify you as the publisher of these copies. The - front cover must present the full title with all words of the title - equally prominent and visible. You may add other material on the - covers in addition. Copying with changes limited to the covers, as - long as they preserve the title of the Document and satisfy these - conditions, can be treated as verbatim copying in other respects. - - If the required texts for either cover are too voluminous to fit - legibly, you should put the first ones listed (as many as fit - reasonably) on the actual cover, and continue the rest onto - adjacent pages. - - If you publish or distribute Opaque copies of the Document - numbering more than 100, you must either include a machine-readable - Transparent copy along with each Opaque copy, or state in or with - each Opaque copy a computer-network location from which the general - network-using public has access to download using public-standard - network protocols a complete Transparent copy of the Document, free - of added material. If you use the latter option, you must take - reasonably prudent steps, when you begin distribution of Opaque - copies in quantity, to ensure that this Transparent copy will - remain thus accessible at the stated location until at least one - year after the last time you distribute an Opaque copy (directly or - through your agents or retailers) of that edition to the public. - - It is requested, but not required, that you contact the authors of - the Document well before redistributing any large number of copies, - to give them a chance to provide you with an updated version of the - Document. - -4. MODIFICATIONS - - You may copy and distribute a Modified Version of the Document - under the conditions of sections 2 and 3 above, provided that you - release the Modified Version under precisely this License, with the - Modified Version filling the role of the Document, thus licensing - distribution and modification of the Modified Version to whoever - possesses a copy of it. In addition, you must do these things in - the Modified Version: - - #+attr_texinfo: :enum A - 1. Use in the Title Page (and on the covers, if any) a title - distinct from that of the Document, and from those of previous - versions (which should, if there were any, be listed in the - History section of the Document). You may use the same title as - a previous version if the original publisher of that version - gives permission. - - 2. List on the Title Page, as authors, one or more persons or - entities responsible for authorship of the modifications in the - Modified Version, together with at least five of the principal - authors of the Document (all of its principal authors, if it has - fewer than five), unless they release you from this requirement. - - 3. State on the Title page the name of the publisher of the - Modified Version, as the publisher. - - 4. Preserve all the copyright notices of the Document. - - 5. Add an appropriate copyright notice for your modifications - adjacent to the other copyright notices. - - 6. Include, immediately after the copyright notices, a license - notice giving the public permission to use the Modified Version - under the terms of this License, in the form shown in the - Addendum below. - - 7. Preserve in that license notice the full lists of Invariant - Sections and required Cover Texts given in the Document's - license notice. - - 8. Include an unaltered copy of this License. - - 9. Preserve the section Entitled "History", Preserve its Title, and - add to it an item stating at least the title, year, new authors, - and publisher of the Modified Version as given on the Title - Page. If there is no section Entitled "History" in the Document, - create one stating the title, year, authors, and publisher of - the Document as given on its Title Page, then add an item - describing the Modified Version as stated in the previous - sentence. - - 10. Preserve the network location, if any, given in the Document - for public access to a Transparent copy of the Document, and - likewise the network locations given in the Document for - previous versions it was based on. These may be placed in the - "History" section. You may omit a network location for a work - that was published at least four years before the Document - itself, or if the original publisher of the version it refers - to gives permission. - - 11. For any section Entitled "Acknowledgements" or "Dedications", - Preserve the Title of the section, and preserve in the section - all the substance and tone of each of the contributor - acknowledgements and/or dedications given therein. - - 12. Preserve all the Invariant Sections of the Document, unaltered - in their text and in their titles. Section numbers or the - equivalent are not considered part of the section titles. - - 13. Delete any section Entitled "Endorsements". Such a section may - not be included in the Modified Version. - - 14. Do not retitle any existing section to be Entitled - "Endorsements" or to conflict in title with any Invariant - Section. - - 15. Preserve any Warranty Disclaimers. - - If the Modified Version includes new front-matter sections or - appendices that qualify as Secondary Sections and contain no material - copied from the Document, you may at your option designate some or all - of these sections as invariant. To do this, add their titles to the - list of Invariant Sections in the Modified Version's license notice. - These titles must be distinct from any other section titles. - - You may add a section Entitled "Endorsements", provided it contains - nothing but endorsements of your Modified Version by various - parties---for example, statements of peer review or that the text has - been approved by an organization as the authoritative definition of a - standard. - - You may add a passage of up to five words as a Front-Cover Text, and a - passage of up to 25 words as a Back-Cover Text, to the end of the list - of Cover Texts in the Modified Version. Only one passage of - Front-Cover Text and one of Back-Cover Text may be added by (or - through arrangements made by) any one entity. If the Document already - includes a cover text for the same cover, previously added by you or - by arrangement made by the same entity you are acting on behalf of, - you may not add another; but you may replace the old one, on explicit - permission from the previous publisher that added the old one. - - The author(s) and publisher(s) of the Document do not by this License - give permission to use their names for publicity for or to assert or - imply endorsement of any Modified Version. - -5. COMBINING DOCUMENTS - - You may combine the Document with other documents released under - this License, under the terms defined in section 4 above for - modified versions, provided that you include in the combination all - of the Invariant Sections of all of the original documents, - unmodified, and list them all as Invariant Sections of your - combined work in its license notice, and that you preserve all - their Warranty Disclaimers. - - The combined work need only contain one copy of this License, and - multiple identical Invariant Sections may be replaced with a single - copy. If there are multiple Invariant Sections with the same name - but different contents, make the title of each such section unique - by adding at the end of it, in parentheses, the name of the - original author or publisher of that section if known, or else - a unique number. Make the same adjustment to the section titles in - the list of Invariant Sections in the license notice of the - combined work. - - In the combination, you must combine any sections Entitled - "History" in the various original documents, forming one section - Entitled "History"; likewise combine any sections Entitled - "Acknowledgements", and any sections Entitled "Dedications". You - must delete all sections Entitled "Endorsements." - -6. COLLECTIONS OF DOCUMENTS - - You may make a collection consisting of the Document and other - documents released under this License, and replace the individual - copies of this License in the various documents with a single copy - that is included in the collection, provided that you follow the - rules of this License for verbatim copying of each of the documents - in all other respects. - - You may extract a single document from such a collection, and - distribute it individually under this License, provided you insert - a copy of this License into the extracted document, and follow this - License in all other respects regarding verbatim copying of that - document. - -7. AGGREGATION WITH INDEPENDENT WORKS - - A compilation of the Document or its derivatives with other - separate and independent documents or works, in or on a volume of - a storage or distribution medium, is called an "aggregate" if the - copyright resulting from the compilation is not used to limit the - legal rights of the compilation's users beyond what the individual - works permit. When the Document is included in an aggregate, this - License does not apply to the other works in the aggregate which - are not themselves derivative works of the Document. - - If the Cover Text requirement of section 3 is applicable to these - copies of the Document, then if the Document is less than one half - of the entire aggregate, the Document's Cover Texts may be placed - on covers that bracket the Document within the aggregate, or the - electronic equivalent of covers if the Document is in electronic - form. Otherwise they must appear on printed covers that bracket - the whole aggregate. - -8. TRANSLATION - - Translation is considered a kind of modification, so you may - distribute translations of the Document under the terms of - section 4. Replacing Invariant Sections with translations requires - special permission from their copyright holders, but you may - include translations of some or all Invariant Sections in addition - to the original versions of these Invariant Sections. You may - include a translation of this License, and all the license notices - in the Document, and any Warranty Disclaimers, provided that you - also include the original English version of this License and the - original versions of those notices and disclaimers. In case of - a disagreement between the translation and the original version of - this License or a notice or disclaimer, the original version will - prevail. - - If a section in the Document is Entitled "Acknowledgements", - "Dedications", or "History", the requirement (section 4) to - Preserve its Title (section 1) will typically require changing the - actual title. - -9. TERMINATION - - You may not copy, modify, sublicense, or distribute the Document - except as expressly provided under this License. Any attempt - otherwise to copy, modify, sublicense, or distribute it is void, - and will automatically terminate your rights under this License. - - However, if you cease all violation of this License, then your - license from a particular copyright holder is reinstated (a) - provisionally, unless and until the copyright holder explicitly and - finally terminates your license, and (b) permanently, if the - copyright holder fails to notify you of the violation by some - reasonable means prior to 60 days after the cessation. - - Moreover, your license from a particular copyright holder is - reinstated permanently if the copyright holder notifies you of the - violation by some reasonable means, this is the first time you have - received notice of violation of this License (for any work) from - that copyright holder, and you cure the violation prior to 30 days - after your receipt of the notice. - - Termination of your rights under this section does not terminate - the licenses of parties who have received copies or rights from you - under this License. If your rights have been terminated and not - permanently reinstated, receipt of a copy of some or all of the - same material does not give you any rights to use it. - -10. FUTURE REVISIONS OF THIS LICENSE - - The Free Software Foundation may publish new, revised versions of - the GNU Free Documentation License from time to time. Such new - versions will be similar in spirit to the present version, but may - differ in detail to address new problems or concerns. See - https://www.gnu.org/copyleft/. - - Each version of the License is given a distinguishing version - number. If the Document specifies that a particular numbered - version of this License "or any later version" applies to it, you - have the option of following the terms and conditions either of - that specified version or of any later version that has been - published (not as a draft) by the Free Software Foundation. If - the Document does not specify a version number of this License, - you may choose any version ever published (not as a draft) by the - Free Software Foundation. If the Document specifies that a proxy - can decide which future versions of this License can be used, that - proxy's public statement of acceptance of a version permanently - authorizes you to choose that version for the Document. - -11. RELICENSING - - "Massive Multiauthor Collaboration Site" (or "MMC Site") means any - World Wide Web server that publishes copyrightable works and also - provides prominent facilities for anybody to edit those works. - A public wiki that anybody can edit is an example of such - a server. A "Massive Multiauthor Collaboration" (or "MMC") - contained in the site means any set of copyrightable works thus - published on the MMC site. - - "CC-BY-SA" means the Creative Commons Attribution-Share Alike 3.0 - license published by Creative Commons Corporation, - a not-for-profit corporation with a principal place of business in - San Francisco, California, as well as future copyleft versions of - that license published by that same organization. - - "Incorporate" means to publish or republish a Document, in whole - or in part, as part of another Document. - - An MMC is "eligible for relicensing" if it is licensed under this - License, and if all works that were first published under this - License somewhere other than this MMC, and subsequently - incorporated in whole or in part into the MMC, (1) had no cover - texts or invariant sections, and (2) were thus incorporated prior - to November 1, 2008. - - The operator of an MMC Site may republish an MMC contained in the - site under CC-BY-SA on the same site at any time before August 1, - 2009, provided the MMC is eligible for relicensing. - -#+texinfo: @page - -* ADDENDUM: How to use this License for your documents -:PROPERTIES: -:UNNUMBERED: notoc -:END: - -To use this License in a document you have written, include a copy of -the License in the document and put the following copyright and -license notices just after the title page: - -#+begin_example - Copyright (C) YEAR YOUR NAME. - Permission is granted to copy, distribute and/or modify this document - under the terms of the GNU Free Documentation License, Version 1.3 - or any later version published by the Free Software Foundation; - with no Invariant Sections, no Front-Cover Texts, and no Back-Cover - Texts. A copy of the license is included in the section entitled ``GNU - Free Documentation License''. -#+end_example - -If you have Invariant Sections, Front-Cover Texts and Back-Cover Texts, -replace the "with...Texts."\nbsp{}line with this: - -#+begin_example - with the Invariant Sections being LIST THEIR TITLES, with - the Front-Cover Texts being LIST, and with the Back-Cover Texts - being LIST. -#+end_example - -If you have Invariant Sections without Cover Texts, or some other -combination of the three, merge those two alternatives to suit the -situation. - -If your document contains nontrivial examples of program code, we -recommend releasing these examples in parallel under your choice of -free software license, such as the GNU General Public License, to -permit their use in free software. diff --git a/docs/res/install-deps.el b/docs/res/install-deps.el deleted file mode 100644 index 09cf4ae..0000000 --- a/docs/res/install-deps.el +++ /dev/null @@ -1,11 +0,0 @@ -(require 'package) -(package-initialize) -(add-to-list 'package-archives '("nongnu" . "https://elpa.nongnu.org/nongnu/") t) -(add-to-list 'package-archives '("melpa" . "https://melpa.org/packages/") t) -(add-to-list 'package-archives '("gnu-devel" . "https://elpa.gnu.org/devel/") t) -(package-refresh-contents) - -(package-install 'org) -(package-install 'org-contrib) -(package-install 'org-transclusion) -(package-install 'htmlize) diff --git a/docs/res/publish.el b/docs/res/publish.el deleted file mode 100644 index c083eb0..0000000 --- a/docs/res/publish.el +++ /dev/null @@ -1,23 +0,0 @@ -(require 'package) -(package-initialize) - -(require 'org) -(require 'org-transclusion) -(require 'ox) -(require 'ox-html) -(require 'htmlize) -(require 'ox-texinfo) -(require 'ox-man) - -(setq org-transclusion-exclude-elements nil - org-html-head-include-default-style nil - org-html-head-include-scripts nil - org-html-postamble-format '(("en" "")) - org-html-postamble t - org-html-divs '((preamble "header" "header") - (content "article" "content") - (postamble "footer" "postamble")) - org-html-doctype "html5" - org-html-htmlize-output-type 'css) - -(org-transclusion-add-all) diff --git a/src/Compiler.v b/src/Compiler.v index cd05aa9..d6d87e2 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -1,22 +1,21 @@ -(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 2019-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 . - *) - (*| +.. + 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 . + ============== Compiler Proof ============== @@ -27,9 +26,12 @@ backwards simulation correct. Imports ======= + We first need to import all of the modules that are used in the correctness proof, which includes all of the passes that are performed in Vericert and the CompCert front end. + +.. coq:: none |*) Require compcert.backend.Selection. @@ -133,24 +135,28 @@ Definition partial_if {A: Type} Definition time {A B: Type} (name: string) (f: A -> B) : A -> B := f. -Definition match_if {A: Type} (flag: unit -> bool) (R: A -> A -> Prop): A -> A -> Prop := +Definition match_if {A: Type} (flag: unit -> bool) (R: A -> A -> Prop) + : A -> A -> Prop := if flag tt then R else eq. Lemma total_if_match: - forall (A: Type) (flag: unit -> bool) (f: A -> A) (rel: A -> A -> Prop) (prog: A), - (forall p, rel p (f p)) -> + forall (A: Type) (flag: unit -> bool) (f: A -> A) + (rel: A -> A -> Prop) (prog: A), + (forall p, rel p (f p)) -> match_if flag rel prog (total_if flag f prog). Proof. intros. unfold match_if, total_if. destruct (flag tt); auto. Qed. Lemma partial_if_match: - forall (A: Type) (flag: unit -> bool) (f: A -> res A) (rel: A -> A -> Prop) (prog tprog: A), + forall (A: Type) (flag: unit -> bool) (f: A -> res A) + (rel: A -> A -> Prop) (prog tprog: A), (forall p tp, f p = OK tp -> rel p tp) -> partial_if flag f prog = OK tprog -> match_if flag rel prog tprog. Proof. - intros. unfold match_if, partial_if in *. destruct (flag tt). auto. congruence. + intros. unfold match_if, partial_if in *. destruct (flag tt). + auto. congruence. Qed. Remark forward_simulation_identity: @@ -164,12 +170,14 @@ Proof. Qed. Lemma match_if_simulation: - forall (A: Type) (sem: A -> semantics) (flag: unit -> bool) (transf: A -> A -> Prop) (prog tprog: A), + forall (A: Type) (sem: A -> semantics) (flag: unit -> bool) + (transf: A -> A -> Prop) (prog tprog: A), match_if flag transf prog tprog -> (forall p tp, transf p tp -> forward_simulation (sem p) (sem tp)) -> forward_simulation (sem prog) (sem tprog). Proof. - intros. unfold match_if in *. destruct (flag tt). eauto. subst. apply forward_simulation_identity. + intros. unfold match_if in *. destruct (flag tt). eauto. subst. + apply forward_simulation_identity. Qed. (*| @@ -188,13 +196,16 @@ Definition transf_backend (r : RTL.program) : res Verilog.program := @@ print (print_RTL 1) @@ Renumber.transf_program @@ print (print_RTL 2) - @@ total_if Compopts.optim_constprop (time "Constant propagation" Constprop.transf_program) + @@ total_if Compopts.optim_constprop + (time "Constant propagation" Constprop.transf_program) @@ print (print_RTL 3) - @@ total_if Compopts.optim_constprop (time "Renumbering" Renumber.transf_program) + @@ total_if Compopts.optim_constprop + (time "Renumbering" Renumber.transf_program) @@ print (print_RTL 4) @@@ partial_if Compopts.optim_CSE (time "CSE" CSE.transf_program) @@ print (print_RTL 5) - @@@ partial_if Compopts.optim_redundancy (time "Redundancy elimination" Deadcode.transf_program) + @@@ partial_if Compopts.optim_redundancy + (time "Redundancy elimination" Deadcode.transf_program) @@ print (print_RTL 6) @@@ time "Unused globals" Unusedglob.transform_program @@ print (print_RTL 7) @@ -220,8 +231,10 @@ Definition transf_hls (p : Csyntax.program) : res Verilog.program := @@ print (print_RTL 0) @@@ transf_backend. -(* This is an unverified version of transf_hls with some experimental additions such as scheduling -that aren't completed yet. *) +(*| +This is an unverified version of transf_hls with some experimental additions +such as scheduling that aren't completed yet. +|*) Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := OK p @@@ SimplExpr.transl_program @@ -234,13 +247,16 @@ Definition transf_hls_temp (p : Csyntax.program) : res Verilog.program := @@ print (print_RTL 1) @@ Renumber.transf_program @@ print (print_RTL 2) - @@ total_if Compopts.optim_constprop (time "Constant propagation" Constprop.transf_program) + @@ total_if Compopts.optim_constprop + (time "Constant propagation" Constprop.transf_program) @@ print (print_RTL 3) - @@ total_if Compopts.optim_constprop (time "Renumbering" Renumber.transf_program) + @@ total_if Compopts.optim_constprop + (time "Renumbering" Renumber.transf_program) @@ print (print_RTL 4) @@@ partial_if Compopts.optim_CSE (time "CSE" CSE.transf_program) @@ print (print_RTL 5) - @@@ partial_if Compopts.optim_redundancy (time "Redundancy elimination" Deadcode.transf_program) + @@@ partial_if Compopts.optim_redundancy + (time "Redundancy elimination" Deadcode.transf_program) @@ print (print_RTL 6) @@@ time "Unused globals" Unusedglob.transform_program @@ print (print_RTL 7) @@ -280,7 +296,8 @@ Definition CompCert's_passes := ::: mkpass (match_if Compopts.optim_CSE CSEproof.match_prog) ::: mkpass (match_if Compopts.optim_redundancy Deadcodeproof.match_prog) ::: mkpass Unusedglobproof.match_prog - ::: (@mkpass _ _ HTLgenproof.match_prog (HTLgenproof.TransfHTLLink HTLgen.transl_program)) + ::: (@mkpass _ _ HTLgenproof.match_prog + (HTLgenproof.TransfHTLLink HTLgen.transl_program)) (*::: mkpass (match_if HLSOpts.optim_ram Memorygen.match_prog)*) ::: mkpass Veriloggenproof.match_prog ::: pass_nil _. @@ -305,22 +322,36 @@ Theorem transf_hls_match: Proof. intros p tp T. unfold transf_hls, time in T. simpl in T. - destruct (SimplExpr.transl_program p) as [p1|e] eqn:P1; simpl in T; try discriminate. - destruct (SimplLocals.transf_program p1) as [p2|e] eqn:P2; simpl in T; try discriminate. - destruct (Cshmgen.transl_program p2) as [p3|e] eqn:P3; simpl in T; try discriminate. - destruct (Cminorgen.transl_program p3) as [p4|e] eqn:P4; simpl in T; try discriminate. - destruct (Selection.sel_program p4) as [p5|e] eqn:P5; simpl in T; try discriminate. + destruct (SimplExpr.transl_program p) as [p1|e] eqn:P1; + simpl in T; try discriminate. + destruct (SimplLocals.transf_program p1) as [p2|e] eqn:P2; + simpl in T; try discriminate. + destruct (Cshmgen.transl_program p2) as [p3|e] eqn:P3; + simpl in T; try discriminate. + destruct (Cminorgen.transl_program p3) as [p4|e] eqn:P4; + simpl in T; try discriminate. + destruct (Selection.sel_program p4) as [p5|e] eqn:P5; + simpl in T; try discriminate. + rewrite ! compose_print_identity in T. + destruct (RTLgen.transl_program p5) as [p6|e] eqn:P6; + simpl in T; try discriminate. + unfold transf_backend, time in T. simpl in T. rewrite ! compose_print_identity in T. - destruct (RTLgen.transl_program p5) as [p6|e] eqn:P6; simpl in T; try discriminate. - unfold transf_backend, time in T. simpl in T. rewrite ! compose_print_identity in T. - destruct (Inlining.transf_program p6) as [p7|e] eqn:P7; simpl in T; try discriminate. + destruct (Inlining.transf_program p6) as [p7|e] eqn:P7; + simpl in T; try discriminate. set (p8 := Renumber.transf_program p7) in *. - set (p9 := total_if Compopts.optim_constprop Constprop.transf_program p8) in *. - set (p10 := total_if Compopts.optim_constprop Renumber.transf_program p9) in *. - destruct (partial_if Compopts.optim_CSE CSE.transf_program p10) as [p11|e] eqn:P11; simpl in T; try discriminate. - destruct (partial_if Compopts.optim_redundancy Deadcode.transf_program p11) as [p12|e] eqn:P12; simpl in T; try discriminate. - destruct (Unusedglob.transform_program p12) as [p13|e] eqn:P13; simpl in T; try discriminate. - destruct (HTLgen.transl_program p13) as [p14|e] eqn:P14; simpl in T; try discriminate. + set (p9 := total_if Compopts.optim_constprop + Constprop.transf_program p8) in *. + set (p10 := total_if Compopts.optim_constprop + Renumber.transf_program p9) in *. + destruct (partial_if Compopts.optim_CSE CSE.transf_program p10) + as [p11|e] eqn:P11; simpl in T; try discriminate. + destruct (partial_if Compopts.optim_redundancy Deadcode.transf_program p11) + as [p12|e] eqn:P12; simpl in T; try discriminate. + destruct (Unusedglob.transform_program p12) as [p13|e] eqn:P13; + simpl in T; try discriminate. + destruct (HTLgen.transl_program p13) as [p14|e] eqn:P14; + simpl in T; try discriminate. (*set (p15 := total_if HLSOpts.optim_ram Memorygen.transf_program p14) in *.*) set (p15 := Veriloggen.transl_program p14) in *. unfold match_prog; simpl. @@ -332,13 +363,18 @@ Proof. exists p6; split. apply RTLgenproof.transf_program_match; auto. exists p7; split. apply Inliningproof.transf_program_match; auto. exists p8; split. apply Renumberproof.transf_program_match; auto. - exists p9; split. apply total_if_match. apply Constpropproof.transf_program_match. - exists p10; split. apply total_if_match. apply Renumberproof.transf_program_match. - exists p11; split. eapply partial_if_match; eauto. apply CSEproof.transf_program_match. - exists p12; split. eapply partial_if_match; eauto. apply Deadcodeproof.transf_program_match. + exists p9; split. apply total_if_match. + apply Constpropproof.transf_program_match. + exists p10; split. apply total_if_match. + apply Renumberproof.transf_program_match. + exists p11; split. eapply partial_if_match; eauto. + apply CSEproof.transf_program_match. + exists p12; split. eapply partial_if_match; eauto. + apply Deadcodeproof.transf_program_match. exists p13; split. apply Unusedglobproof.transf_program_match; auto. exists p14; split. apply HTLgenproof.transf_program_match; auto. - (*exists p15; split. apply total_if_match. apply Memorygen.transf_program_match; auto.*) + (*exists p15; split. apply total_if_match. + apply Memorygen.transf_program_match; auto.*) exists p15; split. apply Veriloggenproof.transf_program_match; auto. inv T. reflexivity. Qed. @@ -347,7 +383,8 @@ Theorem cstrategy_semantic_preservation: forall p tp, match_prog p tp -> forward_simulation (Cstrategy.semantics p) (Verilog.semantics tp) - /\ backward_simulation (atomic (Cstrategy.semantics p)) (Verilog.semantics tp). + /\ backward_simulation (atomic (Cstrategy.semantics p)) + (Verilog.semantics tp). Proof. intros p tp M. unfold match_prog, pass_match in M; simpl in M. Ltac DestructM := @@ -357,7 +394,8 @@ Ltac DestructM := destruct H as (p & M & MM); clear H end. repeat DestructM. subst tp. - assert (F: forward_simulation (Cstrategy.semantics p) (Verilog.semantics p15)). + assert (F: forward_simulation (Cstrategy.semantics p) + (Verilog.semantics p15)). { eapply compose_forward_simulations. eapply SimplExprproof.transl_program_correct; eassumption. @@ -376,24 +414,30 @@ Ltac DestructM := eapply compose_forward_simulations. eapply Renumberproof.transf_program_correct; eassumption. eapply compose_forward_simulations. - eapply match_if_simulation. eassumption. exact Constpropproof.transf_program_correct. + eapply match_if_simulation. eassumption. + exact Constpropproof.transf_program_correct. eapply compose_forward_simulations. - eapply match_if_simulation. eassumption. exact Renumberproof.transf_program_correct. + eapply match_if_simulation. eassumption. + exact Renumberproof.transf_program_correct. eapply compose_forward_simulations. - eapply match_if_simulation. eassumption. exact CSEproof.transf_program_correct. + eapply match_if_simulation. eassumption. + exact CSEproof.transf_program_correct. eapply compose_forward_simulations. - eapply match_if_simulation. eassumption. exact Deadcodeproof.transf_program_correct; eassumption. + eapply match_if_simulation. eassumption. + exact Deadcodeproof.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply Unusedglobproof.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply HTLgenproof.transf_program_correct. eassumption. (*eapply compose_forward_simulations. - eapply match_if_simulation. eassumption. exact Memorygen.transf_program_correct; eassumption.*) + eapply match_if_simulation. eassumption. + exact Memorygen.transf_program_correct; eassumption.*) eapply Veriloggenproof.transf_program_correct; eassumption. } split. auto. apply forward_to_backward_simulation. - apply factor_forward_simulation. auto. eapply sd_traces. eapply Verilog.semantics_determinate. + apply factor_forward_simulation. auto. eapply sd_traces. + eapply Verilog.semantics_determinate. apply atomic_receptive. apply Cstrategy.semantics_strongly_receptive. apply Verilog.semantics_determinate. Qed. @@ -451,11 +495,13 @@ Theorem separate_transf_c_program_correct: link_list c_units = Some c_program -> exists verilog_program, link_list verilog_units = Some verilog_program - /\ backward_simulation (Csem.semantics c_program) (Verilog.semantics verilog_program). + /\ backward_simulation (Csem.semantics c_program) + (Verilog.semantics verilog_program). Proof. intros. assert (nlist_forall2 match_prog c_units verilog_units). - { eapply nlist_forall2_imply. eauto. simpl; intros. apply transf_hls_match; auto. } + { eapply nlist_forall2_imply. eauto. simpl; intros. + apply transf_hls_match; auto. } assert (exists verilog_program, link_list verilog_units = Some verilog_program /\ match_prog c_program verilog_program). { eapply link_list_compose_passes; eauto. } 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 ad61012df574f4e1282049e86d227939ca3d9a9b Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 26 Mar 2022 22:13:58 +0000 Subject: Add theme options --- doc/conf.py | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/doc/conf.py b/doc/conf.py index 015f030..c7a9970 100644 --- a/doc/conf.py +++ b/doc/conf.py @@ -50,6 +50,25 @@ pygments_style = "emacs" # html_theme = 'alabaster' +html_theme_options = { + 'logo': 'images/vericert-main.svg', + 'github_user': 'ymherklotz', + 'github_repo': 'vericert', + 'description': 'A formally verified high-level synthesis tool written in Coq.', + 'github_button': False, + 'show_powered_by': False, +} + +html_sidebars = { + '**': [ + 'about.html', + 'searchbox.html', + 'navigation.html', + 'relations.html', + 'donate.html', + ] +} + # Add any paths that contain custom static files (such as style sheets) here, # relative to this directory. They are copied after the builtin static files, # so a file named "default.css" will overwrite the builtin "default.css". -- cgit From cfa2956933619440ab9803b1468292b191765b38 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 26 Mar 2022 23:03:20 +0000 Subject: Add vericert logo --- doc/_static/images/vericert-main.svg | 500 +++++++++++++++++++++++++++++++++++ 1 file changed, 500 insertions(+) create mode 100644 doc/_static/images/vericert-main.svg diff --git a/doc/_static/images/vericert-main.svg b/doc/_static/images/vericert-main.svg new file mode 100644 index 0000000..2280ec0 --- /dev/null +++ b/doc/_static/images/vericert-main.svg @@ -0,0 +1,500 @@ + +image/svg+xml + + + + + + + + + + + + + + + + + + + + CMYK 80 / 58 / 52 / 55RGB 44 / 60 / 66#2c3c42 + CMYK 45 / 0 / 38 / 0RGB 152 / 217 / 182#98d9b6 + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + b/wcol/wmonochr. + + + + -- cgit From 6959b38a343d4575efc442ea02422dc64cf59d00 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 28 Mar 2022 01:25:28 +0100 Subject: Add to documentation --- default.nix | 1 + doc/Makefile | 4 ++-- doc/_static/css/custom.css | 22 +++++++++++++++++++++- doc/conf.py | 21 ++++----------------- doc/index.rst | 42 +++++++++++++++++++++++++++++++++++++++--- doc/vericert.rst | 8 ++------ 6 files changed, 69 insertions(+), 29 deletions(-) diff --git a/default.nix b/default.nix index eb1125c..341c8fc 100644 --- a/default.nix +++ b/default.nix @@ -15,6 +15,7 @@ stdenv.mkDerivation { ncoqPackages.serapi python3 python3Packages.alectryon + python3Packages.sphinx_rtd_theme ]; enableParallelBuilding = true; diff --git a/doc/Makefile b/doc/Makefile index 5662645..8a34845 100644 --- a/doc/Makefile +++ b/doc/Makefile @@ -9,7 +9,7 @@ SOURCEDIR = . BUILDDIR = _build SOURCE_DATE_EPOCH = $(shell git log -1 --format=%ct) -VS_DOCS := ../src/Compiler.v ../src/hls/RTLBlockInstr.v +VS_DOCS := Compiler.v hls/RTLBlockInstr.v hls/RTLBlockgen.v hls/RTLBlockgenproof.v # Put it first so that "make" without argument is like "make help". help: @@ -20,5 +20,5 @@ help: # Catch-all target: route all unknown targets to Sphinx using the new # "make mode" option. $(O) is meant as a shortcut for $(SPHINXOPTS). %: Makefile - $(foreach d,$(VS_DOCS),cp $(d) $(patsubst ../%,%,$(d));) + $(foreach d,$(VS_DOCS),cp ../src/$(d) src/$(d);) @$(SPHINXBUILD) -M $@ "$(SOURCEDIR)" "$(BUILDDIR)" $(SPHINXOPTS) $(O) diff --git a/doc/_static/css/custom.css b/doc/_static/css/custom.css index 66a1e35..e899c8f 100644 --- a/doc/_static/css/custom.css +++ b/doc/_static/css/custom.css @@ -3,9 +3,29 @@ .alectryon-coqdoc .doc .inlinecode, .alectryon-mref, .alectryon-block, .alectryon-io, -.alectryon-toggle-label, .alectryon-banner, pre, tt, code { +.alectryon-toggle-label, .alectryon-banner, pre, tt, code, +.rst-content kbd, .rst-content pre, .rst-content samp, +.rst-content .linenodiv pre, .rst-content div[class^="highlight"] pre, .rst-content pre.literal-block { font-family: 'Iosevka Fixed Slab', 'Iosevka Slab Web', 'Iosevka Web', 'Iosevka Slab', 'Iosevka', 'Fira Code', monospace; font-feature-settings: "COQX" 1 /* Coq ligatures */, "XV00" 1 /* Legacy */, "calt" 1 /* Fallback */; line-height: initial; + font-size: 0.9em; +} + +img { width: 100%; } +a { color: #469369; } +a:hover { color: #60b386; } + +.wy-menu-vertical header, .wy-menu-vertical p.caption { + color: #fff5db; +} + +.wy-nav-content, .wy-menu-vertical li.current > a, .wy-menu-vertical li.on a { + background-color: #fffbf2; + color: #2a3e40; +} + +.wy-side-nav-search .wy-dropdown > a, .wy-side-nav-search > a { + color: #2a3e40; } diff --git a/doc/conf.py b/doc/conf.py index c7a9970..7000f7b 100644 --- a/doc/conf.py +++ b/doc/conf.py @@ -48,25 +48,12 @@ pygments_style = "emacs" # The theme to use for HTML and HTML Help pages. See the documentation for # a list of builtin themes. # -html_theme = 'alabaster' +html_theme = 'sphinx_rtd_theme' -html_theme_options = { - 'logo': 'images/vericert-main.svg', - 'github_user': 'ymherklotz', - 'github_repo': 'vericert', - 'description': 'A formally verified high-level synthesis tool written in Coq.', - 'github_button': False, - 'show_powered_by': False, -} +html_logo = "_static/images/vericert-main.svg" -html_sidebars = { - '**': [ - 'about.html', - 'searchbox.html', - 'navigation.html', - 'relations.html', - 'donate.html', - ] +html_theme_options = { + 'style_nav_header_background': '#fff5db' } # Add any paths that contain custom static files (such as style sheets) here, diff --git a/doc/index.rst b/doc/index.rst index ebb99df..06aa509 100644 --- a/doc/index.rst +++ b/doc/index.rst @@ -6,21 +6,57 @@ Vericert's Documentation ======================== +A formally verified high-level synthesis (HLS) tool written in Coq, building on top of `CompCert +`_ This ensures the correctness of the C to Verilog translation +according to our Verilog semantics and CompCert's C semantics, removing the need to check the +resulting hardware for behavioural correctness. + +Features +-------- + +The project is currently a work in progress. Currently, the following C features are supported and +have all been proven correct, providing a verified translation from C to Verilog: + +- all int operations, +- non-recursive function calls, +- local arrays and pointers, and +- control-flow structures such as if-statements, for-loops, etc... + .. toctree:: :maxdepth: 2 - :caption: Content: + :caption: Content vericert .. toctree:: :maxdepth: 1 - :caption: Sources: + :caption: Sources src/Compiler src/hls/RTLBlockInstr + src/hls/RTLBlockgen + src/hls/RTLBlockgenproof + +Publications +------------ + +:OOPSLA '21: Yann Herklotz, James D. Pollard, Nadesh Ramanathan, and John Wickerson. Formal + Verification of High-Level Synthesis. In *Proc. ACM Program. Lang.* 5, OOPSLA, 2021. + +:LATTE '21: Yann Herklotz and John Wickerson. High-level synthesis tools should be proven + correct. In *Workshop on Languages, Tools, and Techniques for Accelerator + Design*, 2021. + +Mailing lists +------------- + +For discussions, you can join the following mailing list: `lists.sr.ht/~ymherklotz/vericert-discuss `_. + +For contributing patches to the `sourcehut `_ repository: +`lists.sr.ht/~ymherklotz/vericert-devel `_. Indices -================== +======= * :ref:`genindex` * :ref:`search` diff --git a/doc/vericert.rst b/doc/vericert.rst index 3c8eaeb..19ac28f 100644 --- a/doc/vericert.rst +++ b/doc/vericert.rst @@ -12,7 +12,7 @@ application-specific integrated circuit (ASIC). .. _fig:design: -.. figure:: /_static/images/toolflow.png +.. figure:: /_static/images/toolflow.svg Current design of Vericert, where HTL is an intermediate language representing a finite state machine with data-path (FSMD) and Verilog is the target language. @@ -20,8 +20,6 @@ application-specific integrated circuit (ASIC). The design shown in Figure `fig:design`_ shows how Vericert leverages an existing verified C compiler called `CompCert `_ to perform this translation. -.. index:: vericert - .. _building: Building Vericert @@ -48,8 +46,6 @@ Or by running the test suite using the following command: .. _using-vericert: -.. index:: vericert - Using Vericert -------------- @@ -454,7 +450,7 @@ Example A small standalone Coq file that exhibits many of the style points. -.. code:: coq +.. coq:: no-out (* * Vericert: Verified high-level synthesis. -- 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/common/DecEq.v | 150 +++++++++++++++++++++++++++++++++++++++ src/hls/RTLBlockInstr.v | 22 ++++++ src/hls/RTLBlockgen.v | 171 +++++++++------------------------------------ src/hls/RTLBlockgenproof.v | 77 +++++++++++++------- 4 files changed, 258 insertions(+), 162 deletions(-) create mode 100644 src/common/DecEq.v diff --git a/src/common/DecEq.v b/src/common/DecEq.v new file mode 100644 index 0000000..e9d716f --- /dev/null +++ b/src/common/DecEq.v @@ -0,0 +1,150 @@ +(*| +.. + 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 . + +==================== +Decidable Equalities +==================== +|*) + +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. + +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. 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(-) 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(-) 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(-) 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