aboutsummaryrefslogtreecommitdiffstats
path: root/src/hls/HTL.v
diff options
context:
space:
mode:
authorMichalis Pardalos <m.pardalos@gmail.com>2021-08-26 19:27:17 +0100
committerMichalis Pardalos <m.pardalos@gmail.com>2021-08-26 19:27:17 +0100
commitb067e5c7b0ecce9ffbf21e9c4a7ff96328dec2ba (patch)
treed36a28e84827fe0bb8d7b87d9afa8475b7b2008e /src/hls/HTL.v
parent4dca6dbd48d3e71eb550297d3325555f5e039bb8 (diff)
parent0c021173b3efb1310370de4b2a6f5444c745022f (diff)
downloadvericert-b067e5c7b0ecce9ffbf21e9c4a7ff96328dec2ba.tar.gz
vericert-b067e5c7b0ecce9ffbf21e9c4a7ff96328dec2ba.zip
Merge branch 'oopsla21' into sharing-merge
Diffstat (limited to 'src/hls/HTL.v')
-rw-r--r--src/hls/HTL.v182
1 files changed, 178 insertions, 4 deletions
diff --git a/src/hls/HTL.v b/src/hls/HTL.v
index d777d66..886f86d 100644
--- a/src/hls/HTL.v
+++ b/src/hls/HTL.v
@@ -18,12 +18,13 @@
*)
Require Import Coq.FSets.FMapPositive.
+Require Import Coq.micromega.Lia.
Require compcert.common.Events.
Require compcert.common.Globalenvs.
Require compcert.common.Smallstep.
Require compcert.common.Values.
-Require compcert.lib.Integers.
+Require Import compcert.lib.Integers.
Require Import compcert.lib.Maps.
Require Import vericert.common.Vericertlib.
@@ -33,6 +34,8 @@ Require Import vericert.hls.Array.
Require Import vericert.common.Maps.
Require vericert.hls.Verilog.
+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
@@ -53,7 +56,25 @@ Definition controllogic := PTree.t control_stmnt.
Definition map_well_formed {A : Type} (m : PTree.t A) : Prop :=
forall p0 : positive,
In p0 (map fst (Maps.PTree.elements m)) ->
- Z.pos p0 <= Integers.Int.max_unsigned.
+ (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.
Inductive controlsignal : Type :=
| ctrl_finish : controlsignal
@@ -90,6 +111,11 @@ Record module: Type :=
These will be mapped to the same verilog register. *)
mod_externctrl : AssocMap.t (ident * controlsignal);
mod_wf : (map_well_formed mod_controllogic /\ map_well_formed mod_datapath);
+ 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;
}.
Definition fundef := AST.fundef module.
@@ -161,12 +187,47 @@ Inductive state : Type :=
(m : module)
(args : list value), state.
+Inductive exec_ram:
+ Verilog.reg_associations -> Verilog.arr_associations -> option 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_Some_write:
+ forall ra ar r d_in addr en wr_en u_en,
+ Int.eq en u_en = false ->
+ Int.eq wr_en (ZToValue 0) = false ->
+ (Verilog.assoc_blocking ra)#(ram_en r, 32) = en ->
+ (Verilog.assoc_blocking ra)!(ram_u_en r) = Some u_en ->
+ (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)
+ (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,
+ Int.eq en u_en = false ->
+ (Verilog.assoc_blocking ra)#(ram_en r, 32) = en ->
+ (Verilog.assoc_blocking ra)!(ram_u_en r) = Some u_en ->
+ (Verilog.assoc_blocking ra)!(ram_wr_en r) = Some (ZToValue 0) ->
+ (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.
+
Inductive step : genv -> state -> Events.trace -> state -> Prop :=
| step_module :
forall g mid m st sf ctrl_stmnt data_stmnt
asr asa
basr1 basa1 nasr1 nasa1
basr2 basa2 nasr2 nasa2
+ basr3 basa3 nasr3 nasa3
asr' asa'
f pstval,
asr!(mod_reset m) = Some (ZToValue 0) ->
@@ -187,8 +248,14 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop :=
data_stmnt
(Verilog.mkassociations basr2 nasr2)
(Verilog.mkassociations basa2 nasa2) ->
- asr' = Verilog.merge_regs nasr2 basr2 ->
- asa' = Verilog.merge_arrs nasa2 basa2 ->
+ exec_ram
+ (Verilog.mkassociations (Verilog.merge_regs nasr2 basr2) empty_assocmap)
+ (Verilog.mkassociations (Verilog.merge_arrs nasa2 basa2) (empty_stack m))
+ (mod_ram m)
+ (Verilog.mkassociations basr3 nasr3)
+ (Verilog.mkassociations basa3 nasa3) ->
+ asr' = Verilog.merge_regs nasr3 basr3 ->
+ asa' = Verilog.merge_arrs nasa3 basa3 ->
asr'!(m.(mod_st)) = Some (posToValue pstval) ->
Z.pos pstval <= Integers.Int.max_unsigned ->
step g
@@ -263,3 +330,110 @@ Inductive final_state : state -> Integers.int -> Prop :=
Definition semantics (m : program) :=
Smallstep.Semantics step (initial_state m) final_state
(Globalenvs.Genv.globalenv m).
+
+Definition max_pc_function (m: module) :=
+ List.fold_left Pos.max (List.map fst (PTree.elements m.(mod_controllogic))) 1.
+
+Definition max_list := fold_right Pos.max 1.
+
+Definition max_stmnt_tree t :=
+ PTree.fold (fun i _ st => Pos.max (Verilog.max_reg_stmnt st) i) t 1.
+
+Definition max_reg_ram r :=
+ match r with
+ | None => 1
+ | Some ram => Pos.max (ram_mem ram)
+ (Pos.max (ram_en ram)
+ (Pos.max (ram_addr ram)
+ (Pos.max (ram_addr ram)
+ (Pos.max (ram_wr_en ram)
+ (Pos.max (ram_d_in ram)
+ (Pos.max (ram_d_out ram) (ram_u_en ram)))))))
+ end.
+
+Definition max_reg_module 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))))))))))).
+
+Lemma max_fold_lt :
+ forall m l n, m <= n -> m <= (fold_left Pos.max l n).
+Proof. induction l; crush; apply IHl; lia. Qed.
+
+Lemma max_fold_lt2 :
+ forall (l: list (positive * Verilog.stmnt)) v n,
+ v <= n ->
+ v <= fold_left (fun a p => Pos.max (Verilog.max_reg_stmnt (snd p)) a) l n.
+Proof. induction l; crush; apply IHl; lia. Qed.
+
+Lemma max_fold_lt3 :
+ forall (l: list (positive * Verilog.stmnt)) v v',
+ v <= v' ->
+ fold_left (fun a0 p => Pos.max (Verilog.max_reg_stmnt (snd p)) a0) l v
+ <= fold_left (fun a0 p => Pos.max (Verilog.max_reg_stmnt (snd p)) a0) l v'.
+Proof. induction l; crush; apply IHl; lia. Qed.
+
+Lemma max_fold_lt4 :
+ forall (l: list (positive * Verilog.stmnt)) (a: positive * Verilog.stmnt),
+ fold_left (fun a0 p => Pos.max (Verilog.max_reg_stmnt (snd p)) a0) l 1
+ <= fold_left (fun a0 p => Pos.max (Verilog.max_reg_stmnt (snd p)) a0) l
+ (Pos.max (Verilog.max_reg_stmnt (snd a)) 1).
+Proof. intros; apply max_fold_lt3; lia. Qed.
+
+Lemma max_reg_stmnt_lt_stmnt_tree':
+ forall l (i: positive) v,
+ In (i, v) l ->
+ list_norepet (map fst l) ->
+ Verilog.max_reg_stmnt v <= fold_left (fun a p => Pos.max (Verilog.max_reg_stmnt (snd p)) a) l 1.
+Proof.
+ induction l; crush. inv H; inv H0; simplify. apply max_fold_lt2. lia.
+ transitivity (fold_left (fun (a : positive) (p : positive * Verilog.stmnt) =>
+ Pos.max (Verilog.max_reg_stmnt (snd p)) a) l 1).
+ eapply IHl; eauto. apply max_fold_lt4.
+Qed.
+
+Lemma max_reg_stmnt_le_stmnt_tree :
+ forall d i v,
+ d ! i = Some v ->
+ Verilog.max_reg_stmnt v <= max_stmnt_tree d.
+Proof.
+ intros. unfold max_stmnt_tree. rewrite PTree.fold_spec.
+ apply PTree.elements_correct in H.
+ eapply max_reg_stmnt_lt_stmnt_tree'; eauto.
+ apply PTree.elements_keys_norepet.
+Qed.
+
+Lemma max_reg_stmnt_lt_stmnt_tree :
+ forall d i v,
+ d ! i = Some v ->
+ Verilog.max_reg_stmnt v < max_stmnt_tree d + 1.
+Proof. intros. apply max_reg_stmnt_le_stmnt_tree in H; lia. Qed.
+
+Lemma max_stmnt_lt_module :
+ forall m d i,
+ (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.
+
+Lemma max_list_correct l st : st > max_list l -> Forall (Pos.gt st) l.
+Proof. induction l; crush; constructor; [|apply IHl]; lia. Qed.
+
+Definition max_list_dec (l: list reg) (st: reg) : {Forall (Pos.gt st) l} + {True}.
+ refine (
+ match bool_dec (max_list l <? st) true with
+ | left _ => left _
+ | _ => _
+ end
+ ); auto.
+ apply max_list_correct. apply Pos.ltb_lt in e. lia.
+Qed.