From fdf5e7378715ba1bd8552e2f005c4a15c834a038 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 22 Apr 2020 22:43:03 +0100 Subject: Add stmnt_runp inductive --- src/verilog/Verilog.v | 133 ++++++++++++++++++++++++++++++++++++++++---------- 1 file changed, 106 insertions(+), 27 deletions(-) diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 4713c36..4ad297b 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -173,7 +173,7 @@ Definition fext := PositiveMap.t value. Definition fextclk := nat -> fext. Inductive state: Type := - State: + Runstate: forall (assoc : assoclist) (nbassoc : assoclist) (f : fextclk) @@ -219,7 +219,7 @@ Inductive expr_runp : state -> expr -> value -> Prop := | erun_Vvar : forall assoc na f c v r, assoc!r = Some v -> - expr_runp (State assoc na f c) (Vvar r) v + expr_runp (Runstate assoc na f c) (Vvar r) v | erun_Vinputvar : forall s r v, expr_runp s (Vinputvar r) v @@ -265,7 +265,7 @@ Local Open Scope error_monad_scope. Definition access_fext (s : state) (r : reg) : res value := match s with - | State _ _ f c => + | Runstate _ _ f c => match PositiveMap.find r (f (Pos.to_nat c)) with | Some v => OK v | _ => OK (ZToValue 1 0) @@ -278,7 +278,7 @@ Fixpoint expr_run (s : state) (e : expr) match e with | Vlit v => OK v | Vvar r => match s with - | State assoc _ _ _ => handle_def (ZToValue 32 0) assoc!r + | Runstate assoc _ _ _ => handle_def (ZToValue 32 0) assoc!r end | Vinputvar r => access_fext s r | Vbinop op l r => @@ -307,6 +307,15 @@ Definition assign_name (e : expr) : res reg := | _ => 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 @@ -345,34 +354,83 @@ Fixpoint stmnt_run' (n : nat) (s : state) (st : stmnt) {struct n} : res state := end | Vblock lhs rhs => match s with - | State assoc nbassoc f c => + | Runstate assoc nbassoc f c => do name <- assign_name lhs; do rhse <- expr_run s rhs; - OK (State (PositiveMap.add name rhse assoc) nbassoc f c) + OK (Runstate (PositiveMap.add name rhse assoc) nbassoc f c) end | Vnonblock lhs rhs => match s with - | State assoc nbassoc f c => + | Runstate assoc nbassoc f c => do name <- assign_name lhs; do rhse <- expr_run s rhs; - OK (State assoc (PositiveMap.add name rhse nbassoc) f c) + OK (Runstate assoc (PositiveMap.add name rhse nbassoc) f c) end end | _ => OK s 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. - Definition stmnt_run (s : state) (st : stmnt) : res state := stmnt_run' (stmnt_height st) s st. +Inductive not_matched (s : state) (caseval : value) (curr : expr * stmnt) : Prop := + not_in: + forall currval, + expr_runp s (fst curr) currval -> + caseval <> currval -> + not_matched s caseval curr. + +Inductive stmnt_runp: state -> stmnt -> state -> Prop := +| stmnt_run_Vskip: + forall s, stmnt_runp s Vskip s +| stmnt_run_Vseq: + forall st1 st2 s0 s1 s2, + stmnt_runp s0 st1 s1 -> + stmnt_runp s1 st2 s2 -> + stmnt_runp s0 (Vseq st1 st2) s2 +| stmnt_runp_Vcond_true: + forall s0 s1 c vc stt stf, + expr_runp s0 c vc -> + valueToBool vc = true -> + stmnt_runp s0 stt s1 -> + stmnt_runp s0 (Vcond c stt stf) s1 +| stmnt_runp_Vcond_false: + forall s0 s1 c vc stt stf, + expr_runp s0 c vc -> + valueToBool vc = false -> + stmnt_runp s0 stf s1 -> + stmnt_runp s0 (Vcond c stt stf) s1 +| stmnt_runp_Vcase: + forall e ve s0 s1 me mve sc clst def, + expr_runp s0 e ve -> + expr_runp s0 me mve -> + mve = ve -> + In (me, sc) clst -> + stmnt_runp s0 sc s1 -> + stmnt_runp s0 (Vcase e clst def) s1 +| stmnt_runp_Vcase_default: + forall s0 st s1 e ve clst, + expr_runp s0 e ve -> + Forall (not_matched s0 ve) clst -> + stmnt_runp s0 st s1 -> + stmnt_runp s0 (Vcase e clst (Some st)) s1 +| stmnt_runp_Vblock: + forall lhs name rhs rhsval s assoc assoc' nbassoc f cycle, + assign_name lhs = OK name -> + expr_runp s rhs rhsval -> + assoc' = (PositiveMap.add name rhsval assoc) -> + stmnt_runp (Runstate assoc nbassoc f cycle) + (Vblock lhs rhs) + (Runstate assoc' nbassoc f cycle) +| stmnt_runp_Vnonblock: + forall lhs name rhs rhsval s assoc nbassoc nbassoc' f cycle, + assign_name lhs = OK name -> + expr_runp s rhs rhsval -> + nbassoc' = (PositiveMap.add name rhsval nbassoc) -> + stmnt_runp (Runstate assoc nbassoc f cycle) + (Vblock lhs rhs) + (Runstate assoc nbassoc' f cycle). + Fixpoint mi_step (s : state) (m : list module_item) {struct m} : res state := let run := fun st ls => do s' <- stmnt_run s st; @@ -396,8 +454,8 @@ Definition empty_assoclist : assoclist := PositiveMap.empty value. Definition mi_step_commit (s : state) (m : list module_item) : res state := match mi_step s m with - | OK (State assoc nbassoc f c) => - OK (State (merge_assoclist nbassoc assoc) empty_assoclist f c) + | OK (Runstate assoc nbassoc f c) => + OK (Runstate (merge_assoclist nbassoc assoc) empty_assoclist f c) | Error msg => Error msg end. @@ -406,19 +464,25 @@ Fixpoint mi_run (f : fextclk) (assoc : assoclist) (m : list module_item) (n : na match n with | S n' => do assoc' <- mi_run f assoc m n'; - match mi_step_commit (State assoc' empty_assoclist f (Pos.of_nat n')) m with - | OK (State assoc _ _ _) => OK assoc + match mi_step_commit (Runstate assoc' empty_assoclist f (Pos.of_nat n')) m with + | OK (Runstate 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 => (add_assoclist (fst (mod_reset m)) (ZToValue 1 1) empty_assoclist) + | _ => (add_assoclist (fst (mod_reset m)) (ZToValue 1 0) empty_assoclist) + end. + Definition module_run (n : nat) (m : module) : res assoclist := - let f := fun x => match x with - | S O => (add_assoclist (fst (mod_reset m)) (ZToValue 1 1) empty_assoclist) - | _ => (add_assoclist (fst (mod_reset m)) (ZToValue 1 0) empty_assoclist) - end in - mi_run f empty_assoclist (mod_body m) n. + mi_run (initial_fextclk m) empty_assoclist (mod_body m) n. Local Close Scope error_monad_scope. @@ -459,4 +523,19 @@ Proof. assumption. Qed. -*) + *) + +Inductive initial_state (m: module): state -> Prop := +| initial_state_intro: + initial_state m (Runstate empty_assoclist empty_assoclist (initial_fextclk m) xH). + +(** A final state is a [Returnstate] with an empty call stack. *) + +Inductive final_state: state -> int -> Prop := + | final_state_intro: forall r m, + final_state (Returnstate nil (Vint r) m) r. + +(** The small-step semantics for a module. *) + +Definition semantics (p: module) := + Semantics step (initial_state p) final_state (Genv.globalenv p). -- cgit From 6557bd4ad48626d1c7f644f4134560d363786766 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 23 Apr 2020 22:40:18 +0100 Subject: Add OS detection to makefile --- Makefile | 21 +++++++++++++-------- README.md | 8 -------- 2 files changed, 13 insertions(+), 16 deletions(-) diff --git a/Makefile b/Makefile index 1d328ec..42e6a45 100644 --- a/Makefile +++ b/Makefile @@ -1,18 +1,23 @@ +UNAME_S := $(shell uname -s) +ifeq ($(UNAME_S),Linux) + ARCH := x86_64-linux +endif +ifeq ($(UNAME_S),Darwin) + ARCH := x86_64-macosx +endif + COMPCERTRECDIRS := lib common x86_64 x86 backend cfrontend driver flocq exportclight \ MenhirLib cparser -COMPCERTCOQINCLUDES := $(foreach d, $(COMPCERTRECDIRS), -R lib/CompCert/$(d) compcert.$(d)) - COQINCLUDES := -R src/common coqup.common -R src/verilog coqup.verilog \ -R src/extraction coqup.extraction -R src/translation coqup.translation \ - -R src coqup $(COMPCERTCOQINCLUDES) + -R src coqup \ + $(foreach d, $(COMPCERTRECDIRS), -R lib/CompCert/$(d) compcert.$(d)) COQEXEC := $(COQBIN)coqtop $(COQINCLUDES) -batch -load-vernac-source -COQMAKE := "$(COQBIN)coq_makefile" +COQMAKE := $(COQBIN)coq_makefile -COQUPDIRS := translation common verilog -VSSUBDIR := $(foreach d, $(COQUPDIRS), src/$(d)/*.v) -VS := src/Compiler.v src/Simulator.v $(VSSUBDIR) +VS := src/Compiler.v src/Simulator.v $(foreach d, translation common verilog, src/$(d)/*.v) PREFIX ?= . @@ -23,7 +28,7 @@ all: lib/COMPCERTSTAMP $(MAKE) compile lib/COMPCERTSTAMP: - (cd lib/CompCert && ./configure x86_64-linux) + (cd lib/CompCert && ./configure $(ARCH)) $(MAKE) -C lib/CompCert touch $@ diff --git a/README.md b/README.md index df65f8c..026266b 100644 --- a/README.md +++ b/README.md @@ -18,14 +18,6 @@ The project is written in Coq, a theorem prover, which is extracted to OCaml so These dependencies can be installed manually, or automatically through Nix. -### Building on OSX - -To build the project on OSX, currently the makefile has to be manually edited so that CompCert builds for the correct architecture. To do this, simply execute the following `sed` command to change the instance of `x86_64-linux` into `x86_64-macosx`. - -``` shell -sed -i'' 's/x86_64-linux/x86_64-macosx/' Makefile -``` - ### Downloading 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, you can run: -- cgit From 57d74093a19569b9dd55520c6a1548e4827f3b1e Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 24 Apr 2020 22:16:01 +0100 Subject: Add valueToInt function --- src/verilog/Value.v | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index fe53dbc..241534c 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -78,6 +78,9 @@ Definition uvalueToZ (v : value) : Z := Definition intToValue (i : Integers.int) : value := ZToValue Int.wordsize (Int.unsigned i). +Definition valueToInt (i : value) : Integers.int := + Int.repr (valueToZ i). + (** 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 -- cgit From eee95c4fe9f8b0e1d24224f9c980c16156d9b80d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 24 Apr 2020 22:16:15 +0100 Subject: Add CompCert semantics for Verilog --- src/verilog/Verilog.v | 233 ++++++++++++++++++++++++++++++++------------------ 1 file changed, 152 insertions(+), 81 deletions(-) diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 4ad297b..a4504c2 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -28,9 +28,8 @@ From Coq Require Import Import ListNotations. From coqup Require Import common.Coquplib common.Show verilog.Value. - -From compcert Require Integers. -From compcert Require Import Errors. +From compcert Require Integers Events. +From compcert Require Import Errors Smallstep Globalenvs. Notation "a ! b" := (PositiveMap.find b a) (at level 1). @@ -157,6 +156,7 @@ Record module : Type := mkmodule { mod_clk : szreg; mod_finish : szreg; mod_return : szreg; + mod_state : szreg; (**r Variable that defines the current state, it should be internal. *) mod_args : list szreg; mod_body : list module_item; }. @@ -172,12 +172,40 @@ Coercion Vvar : reg >-> expr. Definition fext := PositiveMap.t value. Definition fextclk := nat -> fext. -Inductive state: Type := - Runstate: - forall (assoc : assoclist) +(** ** State + +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. +- 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. *) + +Inductive state : Type := +| Runstate: + forall (m : module) + (mi : module_item) + (mis : list module_item) + (assoc : assoclist) (nbassoc : assoclist) (f : fextclk) - (cycle : positive), + (cycle : nat) + (pc : positive), + state +| Finishstate: + forall v : value, state. Definition binop_run (op : binop) : forall v1 v2 : value, vsize v1 = vsize v2 -> value := @@ -212,40 +240,41 @@ Definition unop_run (op : unop) : value -> value := | Vnot => vbitneg end. -Inductive expr_runp : state -> expr -> value -> Prop := +Inductive expr_runp : fext -> assoclist -> expr -> value -> Prop := | erun_Vlit : - forall s v, - expr_runp s (Vlit v) v + forall fext assoc v, + expr_runp fext assoc (Vlit v) v | erun_Vvar : - forall assoc na f c v r, + forall fext assoc v r, assoc!r = Some v -> - expr_runp (Runstate assoc na f c) (Vvar r) v + expr_runp fext assoc (Vvar r) v | erun_Vinputvar : - forall s r v, - expr_runp s (Vinputvar r) v + forall fext assoc r v, + fext!r = Some v -> + expr_runp fext assoc (Vinputvar r) v | erun_Vbinop : - forall s op l r lv rv oper EQ, - expr_runp s l lv -> - expr_runp s r rv -> + forall fext assoc op l r lv rv oper EQ, + expr_runp fext assoc l lv -> + expr_runp fext assoc r rv -> oper = binop_run op -> - expr_runp s (Vbinop op l r) (oper lv rv EQ) + expr_runp fext assoc (Vbinop op l r) (oper lv rv EQ) | erun_Vunop : - forall s u vu op oper, - expr_runp s u vu -> + forall fext assoc u vu op oper, + expr_runp fext assoc u vu -> oper = unop_run op -> - expr_runp s (Vunop op u) (oper vu) + expr_runp fext assoc (Vunop op u) (oper vu) | erun_Vternary_true : - forall s c ts fs vc vt, - expr_runp s c vc -> - expr_runp s ts vt -> + forall fext assoc c ts fs vc vt, + expr_runp fext assoc c vc -> + expr_runp fext assoc ts vt -> valueToBool vc = true -> - expr_runp s (Vternary c ts fs) vt + expr_runp fext assoc (Vternary c ts fs) vt | erun_Vternary_false : - forall s c ts fs vc vf, - expr_runp s c vc -> - expr_runp s fs vf -> + forall fext assoc c ts fs vc vf, + expr_runp fext assoc c vc -> + expr_runp fext assoc fs vf -> valueToBool vc = false -> - expr_runp s (Vternary c ts fs) vf. + expr_runp fext assoc (Vternary c ts fs) vf. Definition handle_opt {A : Type} (err : errmsg) (val : option A) : res A := @@ -263,22 +292,20 @@ Definition handle_def {A : Type} (a : A) (val : option A) Local Open Scope error_monad_scope. -Definition access_fext (s : state) (r : reg) : res value := - match s with - | Runstate _ _ f c => - match PositiveMap.find r (f (Pos.to_nat c)) with - | Some v => OK v - | _ => OK (ZToValue 1 0) - end +Definition access_fext (f : fext) (r : reg) : res value := + match PositiveMap.find r f with + | Some v => OK v + | _ => OK (ZToValue 1 0) end. (* TODO FIX Vvar case without default *) -Fixpoint expr_run (s : state) (e : expr) +(*Fixpoint expr_run (assoc : assoclist) (e : expr) {struct e} : res value := match e with | Vlit v => OK v | Vvar r => match s with - | Runstate assoc _ _ _ => handle_def (ZToValue 32 0) assoc!r + | Runstate _ assoc _ _ _ => handle_def (ZToValue 32 0) assoc!r + | _ => Error (msg "Verilog: Wrong state") end | Vinputvar r => access_fext s r | Vbinop op l r => @@ -296,7 +323,7 @@ Fixpoint expr_run (s : state) (e : expr) | Vternary c te fe => do cv <- expr_run s c; if valueToBool cv then expr_run s te else expr_run s fe - end. + end.*) (** Return the name of the lhs of an assignment. For now, this function is quite simple because only assignment to normal variables is supported and needed. *) @@ -307,7 +334,7 @@ Definition assign_name (e : expr) : res reg := | _ => Error (msg "Verilog: expression not supported on lhs of assignment") end. -Fixpoint stmnt_height (st : stmnt) {struct st} : nat := +(*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)) @@ -354,31 +381,38 @@ Fixpoint stmnt_run' (n : nat) (s : state) (st : stmnt) {struct n} : res state := end | Vblock lhs rhs => match s with - | Runstate assoc nbassoc f c => + | Runstate m assoc nbassoc f c => do name <- assign_name lhs; do rhse <- expr_run s rhs; - OK (Runstate (PositiveMap.add name rhse assoc) nbassoc f c) + OK (Runstate m (PositiveMap.add name rhse assoc) nbassoc f c) + | _ => Error (msg "Verilog: Wrong state") end | Vnonblock lhs rhs => match s with - | Runstate assoc nbassoc f c => + | Runstate m assoc nbassoc f c => do name <- assign_name lhs; do rhse <- expr_run s rhs; - OK (Runstate assoc (PositiveMap.add name rhse nbassoc) f c) + OK (Runstate 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. + stmnt_run' (stmnt_height st) s st. *) -Inductive not_matched (s : state) (caseval : value) (curr : expr * stmnt) : Prop := +Inductive not_matched : state -> value -> expr * stmnt -> Prop := not_in: - forall currval, - expr_runp s (fst curr) currval -> + forall caseval curr currval m mi c assoc nbassoc f cycle pc, + expr_runp (f cycle) assoc (fst curr) currval -> caseval <> currval -> - not_matched s caseval curr. + not_matched (Runstate m mi c assoc nbassoc f cycle pc) caseval curr. + +Inductive state_fext_assoc : state -> fext * assoclist -> Prop := +| get_state_fext_assoc : + forall m mi c assoc nbassoc f cycle pc, + state_fext_assoc (Runstate m mi c assoc nbassoc f cycle pc) (f cycle, assoc). Inductive stmnt_runp: state -> stmnt -> state -> Prop := | stmnt_run_Vskip: @@ -389,49 +423,53 @@ Inductive stmnt_runp: state -> stmnt -> state -> Prop := stmnt_runp s1 st2 s2 -> stmnt_runp s0 (Vseq st1 st2) s2 | stmnt_runp_Vcond_true: - forall s0 s1 c vc stt stf, - expr_runp s0 c vc -> + forall s0 f assoc s1 c vc stt stf, + state_fext_assoc s0 (f, assoc) -> + expr_runp f assoc c vc -> valueToBool vc = true -> stmnt_runp s0 stt s1 -> stmnt_runp s0 (Vcond c stt stf) s1 | stmnt_runp_Vcond_false: - forall s0 s1 c vc stt stf, - expr_runp s0 c vc -> + forall s0 f assoc s1 c vc stt stf, + state_fext_assoc s0 (f, assoc) -> + expr_runp f assoc c vc -> valueToBool vc = false -> stmnt_runp s0 stf s1 -> stmnt_runp s0 (Vcond c stt stf) s1 | stmnt_runp_Vcase: - forall e ve s0 s1 me mve sc clst def, - expr_runp s0 e ve -> - expr_runp s0 me mve -> + forall e ve s0 f assoc s1 me mve sc clst def, + state_fext_assoc s0 (f, assoc) -> + expr_runp f assoc e ve -> + expr_runp f assoc me mve -> mve = ve -> In (me, sc) clst -> stmnt_runp s0 sc s1 -> stmnt_runp s0 (Vcase e clst def) s1 | stmnt_runp_Vcase_default: - forall s0 st s1 e ve clst, - expr_runp s0 e ve -> + forall s0 f assoc st s1 e ve clst, + state_fext_assoc s0 (f, assoc) -> + expr_runp f assoc e ve -> Forall (not_matched s0 ve) clst -> stmnt_runp s0 st s1 -> stmnt_runp s0 (Vcase e clst (Some st)) s1 | stmnt_runp_Vblock: - forall lhs name rhs rhsval s assoc assoc' nbassoc f cycle, + forall lhs name rhs rhsval m mi c assoc assoc' nbassoc f cycle pc, assign_name lhs = OK name -> - expr_runp s rhs rhsval -> + expr_runp (f cycle) assoc rhs rhsval -> assoc' = (PositiveMap.add name rhsval assoc) -> - stmnt_runp (Runstate assoc nbassoc f cycle) + stmnt_runp (Runstate m mi c assoc nbassoc f cycle pc) (Vblock lhs rhs) - (Runstate assoc' nbassoc f cycle) + (Runstate m mi c assoc' nbassoc f cycle pc) | stmnt_runp_Vnonblock: - forall lhs name rhs rhsval s assoc nbassoc nbassoc' f cycle, + forall lhs name rhs rhsval m mi c assoc nbassoc nbassoc' f cycle pc, assign_name lhs = OK name -> - expr_runp s rhs rhsval -> + expr_runp (f cycle) assoc rhs rhsval -> nbassoc' = (PositiveMap.add name rhsval nbassoc) -> - stmnt_runp (Runstate assoc nbassoc f cycle) + stmnt_runp (Runstate m mi c assoc nbassoc f cycle pc) (Vblock lhs rhs) - (Runstate assoc nbassoc' f cycle). + (Runstate m mi c assoc nbassoc' f cycle pc). -Fixpoint mi_step (s : state) (m : list module_item) {struct m} : res state := +(*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 @@ -442,7 +480,24 @@ Fixpoint mi_step (s : state) (m : list module_item) {struct m} : res state := | (Valways_comb _ st)::ls => run st ls | (Vdecl _ _)::ls => mi_step s ls | nil => OK s - end. + end.*) + +Inductive mi_stepp : state -> module_item -> state -> Prop := +| mi_stepp_Valways : + forall s0 st s1 c, + stmnt_runp s0 st s1 -> + mi_stepp s0 (Valways c st) s1 +| mi_stepp_Valways_ff : + forall s0 st s1 c, + stmnt_runp s0 st s1 -> + mi_stepp s0 (Valways_ff c st) s1 +| mi_stepp_Valways_comb : + forall s0 st s1 c, + stmnt_runp s0 st s1 -> + mi_stepp s0 (Valways_comb c st) s1 +| mi_stepp_Vdecl : + forall s lhs rhs, + mi_stepp s (Vdecl lhs rhs) s. Definition add_assoclist (r : reg) (v : value) (assoc : assoclist) : assoclist := PositiveMap.add r v assoc. @@ -452,14 +507,15 @@ Definition merge_assoclist (nbassoc assoc : assoclist) : assoclist := Definition empty_assoclist : assoclist := PositiveMap.empty value. -Definition mi_step_commit (s : state) (m : list module_item) : res state := +(*Definition mi_step_commit (s : state) (m : list module_item) : res state := match mi_step s m with - | OK (Runstate assoc nbassoc f c) => - OK (Runstate (merge_assoclist nbassoc assoc) empty_assoclist f c) + | OK (Runstate m assoc nbassoc f c) => + OK (Runstate m (merge_assoclist nbassoc assoc) empty_assoclist f c) | Error msg => Error msg - end. + | _ => Error (msg "Verilog: Wrong state") + end.*) -Fixpoint mi_run (f : fextclk) (assoc : assoclist) (m : list module_item) (n : nat) +(*Fixpoint mi_run (f : fextclk) (assoc : assoclist) (m : list module_item) (n : nat) {struct n} : res assoclist := match n with | S n' => @@ -469,7 +525,7 @@ Fixpoint mi_run (f : fextclk) (assoc : assoclist) (m : list module_item) (n : na | Error m => Error m end | O => OK assoc - end. + 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 @@ -481,8 +537,8 @@ Definition initial_fextclk (m : module) : fextclk := | _ => (add_assoclist (fst (mod_reset m)) (ZToValue 1 0) empty_assoclist) end. -Definition module_run (n : nat) (m : module) : res assoclist := - mi_run (initial_fextclk m) empty_assoclist (mod_body m) n. +(*Definition module_run (n : nat) (m : module) : res assoclist := + mi_run (initial_fextclk m) empty_assoclist (mod_body m) n.*) Local Close Scope error_monad_scope. @@ -525,17 +581,32 @@ Qed. *) +Definition genv := Genv.t unit unit. + +Inductive step : genv -> state -> Events.trace -> state -> Prop := +| step_module : + forall g m mi hmi tmi assoc nbassoc f cycle pc e, + step g (Runstate m mi (hmi::tmi) assoc nbassoc f cycle pc) e + (Runstate m hmi tmi assoc nbassoc f cycle pc) +| step_module_empty : + forall g tmi hmi m mi assoc nbassoc f cycle pc e, + hmi::tmi = mod_body m -> + step g (Runstate m mi nil assoc nbassoc f cycle pc) e + (Runstate m hmi tmi assoc nbassoc f (S cycle) pc). + Inductive initial_state (m: module): state -> Prop := | initial_state_intro: - initial_state m (Runstate empty_assoclist empty_assoclist (initial_fextclk m) xH). + forall hmi tmi, + hmi::tmi = mod_body m -> + initial_state m (Runstate m hmi tmi empty_assoclist empty_assoclist (initial_fextclk m) O xH). (** A final state is a [Returnstate] with an empty call stack. *) -Inductive final_state: state -> int -> Prop := - | final_state_intro: forall r m, - final_state (Returnstate nil (Vint r) m) r. +Inductive final_state: state -> Integers.int -> Prop := + | final_state_intro: forall v, + final_state (Finishstate v) (valueToInt v). (** The small-step semantics for a module. *) Definition semantics (p: module) := - Semantics step (initial_state p) final_state (Genv.globalenv p). + Semantics step (initial_state p) final_state (Genv.empty_genv unit unit nil). -- cgit From 8e012d446357cd4f7f1ef25814ea7891c45086ae Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 29 Apr 2020 18:04:27 +0100 Subject: Add documentation and conform to specification --- src/translation/Veriloggen.v | 65 ++++++++++++++++++++++++++++---------------- 1 file changed, 41 insertions(+), 24 deletions(-) diff --git a/src/translation/Veriloggen.v b/src/translation/Veriloggen.v index 6aa94df..aa12a67 100644 --- a/src/translation/Veriloggen.v +++ b/src/translation/Veriloggen.v @@ -359,6 +359,10 @@ Definition add_instr (n : node) (n' : node) (st : stmnt) : mon unit := | _, _, _ => Error (Errors.msg "Veriloggen.add_instr") end. +(** Add a register to the state without changing the [st_freshreg], as this +should already be the right value. This can be assumed because the max register +is taken from the RTL code. *) + Definition add_reg (r: reg) (s: state) : state := mkstate (st_freshreg s) (st_freshstate s) @@ -398,6 +402,9 @@ Lemma decl_io_state_incr: (st_wf s)). Proof. constructor; simpl; auto using Ple_succ with verilog_state. Qed. +(** Declare a new register by incrementing the [st_freshreg] by one and +returning the old [st_freshreg]. *) + Definition decl_io (sz: nat): mon (reg * nat) := fun s => let r := s.(st_freshreg) in OK (r, sz) (mkstate @@ -409,6 +416,9 @@ Definition decl_io (sz: nat): mon (reg * nat) := (st_wf s)) (decl_io_state_incr s). +(** Declare a new register which is given, by changing [st_decl] and without +affecting anything else. *) + Definition declare_reg (r: reg) (sz: nat): mon (reg * nat) := fun s => OK (r, sz) (mkstate (st_freshreg s) @@ -419,6 +429,9 @@ Definition declare_reg (r: reg) (sz: nat): mon (reg * nat) := (st_wf s)) (change_decl_state_incr s (PositiveMap.add r sz s.(st_decl))). +(** Declare a new fresh register by first getting a valid register to increment, +and then adding it to the declaration list. *) + Definition decl_fresh_reg (sz : nat) : mon (reg * nat) := do (r, s) <- decl_io sz; declare_reg r s. @@ -529,11 +542,8 @@ Definition make_statetrans_cases (r : reg) (st : positive * statetrans) : expr * Definition make_statetrans (r : reg) (s : PositiveMap.t statetrans) : stmnt := Vcase (Vvar r) (map (make_statetrans_cases r) (PositiveMap.elements s)) (Some Vskip). -Fixpoint allocate_regs (e : list (reg * nat)) {struct e} : list module_item := - match e with - | (r, n)::es => Vdecl r n :: allocate_regs es - | nil => nil - end. +Definition allocate_reg (e : reg * nat) : module_item := + match e with (r, n) => Vdecl r n end. Definition make_module_items (entry: node) (clk st rst: reg) (s: state) : list module_item := (Valways_ff (Vposedge clk) @@ -541,7 +551,7 @@ Definition make_module_items (entry: node) (clk st rst: reg) (s: state) : list m (nonblock st (posToExpr 32 entry)) (make_statetrans st s.(st_statetrans)))) :: (Valways_ff (Vposedge clk) (make_stm st s.(st_stm))) - :: (allocate_regs (PositiveMap.elements s.(st_decl))). + :: (map allocate_reg (PositiveMap.elements s.(st_decl))). (** To start out with, the assumption is made that there is only one function/module in the original code. *) @@ -549,27 +559,16 @@ Definition make_module_items (entry: node) (clk st rst: reg) (s: state) : list m Definition set_int_size (r: reg) : reg * nat := (r, 32%nat). Definition transf_module (f: function) : mon module := - do fin <- decl_io 1%nat; - do rtrn <- decl_io 32%nat; + do fin <- decl_io 1; + do rtrn <- decl_io 32; do _ <- traverselist (transf_instr (fst fin) (fst rtrn)) (Maps.PTree.elements f.(fn_code)); - do start <- decl_io 1%nat; - do rst <- decl_io 1%nat; - do clk <- decl_io 1%nat; - do st <- decl_fresh_reg 32%nat; + do start <- decl_io 1; + do rst <- decl_io 1; + do clk <- decl_io 1; + do st <- decl_fresh_reg 32; do current_state <- get; let mi := make_module_items f.(fn_entrypoint) (fst clk) (fst st) (fst rst) current_state in - ret (mkmodule start rst clk fin rtrn (map set_int_size f.(fn_params)) mi). - -Fixpoint main_function (main : ident) (flist : list (ident * AST.globdef fundef unit)) -{struct flist} : option function := - match flist with - | (i, AST.Gfun (AST.Internal f)) :: xs => - if Pos.eqb i main - then Some f - else main_function main xs - | _ :: xs => main_function main xs - | nil => None - end. + ret (mkmodule start rst clk fin rtrn st (map set_int_size f.(fn_params)) mi). Lemma max_state_valid_freshstate: forall f, @@ -591,6 +590,24 @@ Definition max_state (f: function) : state := (st_decl init_state) (max_state_st_wf f). +Definition transl_module (f : function) : Errors.res module := + run_mon (max_state f) (transf_module f). + +(** Retrieve the main function from the RTL, this should probably be replaced by +retrieving a designated function instead. This could be passed on the +commandline and be assumed as a [Parameter] in this code. *) + +Fixpoint main_function (main : ident) (flist : list (ident * AST.globdef fundef unit)) +{struct flist} : option function := + match flist with + | (i, AST.Gfun (AST.Internal f)) :: xs => + if Pos.eqb i main + then Some f + else main_function main xs + | _ :: xs => main_function main xs + | nil => None + end. + Definition transf_program (d : program) : Errors.res module := match main_function d.(AST.prog_main) d.(AST.prog_defs) with | Some f => run_mon (max_state f) (transf_module f) -- cgit From 8f370a52a5424389d9f0b4b044020932a2febfd4 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 3 May 2020 17:12:09 +0100 Subject: Change to State --- src/verilog/Verilog.v | 43 ++++++++++++++++++++++--------------------- 1 file changed, 22 insertions(+), 21 deletions(-) diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index a4504c2..7a3982c 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -194,7 +194,7 @@ retrieved and set appropriately. of the Verilog that should match with the RTL. *) Inductive state : Type := -| Runstate: +| State: forall (m : module) (mi : module_item) (mis : list module_item) @@ -304,7 +304,7 @@ Definition access_fext (f : fext) (r : reg) : res value := match e with | Vlit v => OK v | Vvar r => match s with - | Runstate _ assoc _ _ _ => handle_def (ZToValue 32 0) assoc!r + | State _ assoc _ _ _ => handle_def (ZToValue 32 0) assoc!r | _ => Error (msg "Verilog: Wrong state") end | Vinputvar r => access_fext s r @@ -381,18 +381,18 @@ Fixpoint stmnt_run' (n : nat) (s : state) (st : stmnt) {struct n} : res state := end | Vblock lhs rhs => match s with - | Runstate m assoc nbassoc f c => + | State m assoc nbassoc f c => do name <- assign_name lhs; do rhse <- expr_run s rhs; - OK (Runstate m (PositiveMap.add name rhse assoc) nbassoc f c) + OK (State m (PositiveMap.add name rhse assoc) nbassoc f c) | _ => Error (msg "Verilog: Wrong state") end | Vnonblock lhs rhs => match s with - | Runstate m assoc nbassoc f c => + | State m assoc nbassoc f c => do name <- assign_name lhs; do rhse <- expr_run s rhs; - OK (Runstate m assoc (PositiveMap.add name rhse nbassoc) f c) + OK (State m assoc (PositiveMap.add name rhse nbassoc) f c) | _ => Error (msg "Verilog: Wrong state") end end @@ -407,12 +407,12 @@ Inductive not_matched : state -> value -> expr * stmnt -> Prop := forall caseval curr currval m mi c assoc nbassoc f cycle pc, expr_runp (f cycle) assoc (fst curr) currval -> caseval <> currval -> - not_matched (Runstate m mi c assoc nbassoc f cycle pc) caseval curr. + not_matched (State m mi c assoc nbassoc f cycle pc) caseval curr. Inductive state_fext_assoc : state -> fext * assoclist -> Prop := | get_state_fext_assoc : forall m mi c assoc nbassoc f cycle pc, - state_fext_assoc (Runstate m mi c assoc nbassoc f cycle pc) (f cycle, assoc). + state_fext_assoc (State m mi c assoc nbassoc f cycle pc) (f cycle, assoc). Inductive stmnt_runp: state -> stmnt -> state -> Prop := | stmnt_run_Vskip: @@ -457,17 +457,17 @@ Inductive stmnt_runp: state -> stmnt -> state -> Prop := assign_name lhs = OK name -> expr_runp (f cycle) assoc rhs rhsval -> assoc' = (PositiveMap.add name rhsval assoc) -> - stmnt_runp (Runstate m mi c assoc nbassoc f cycle pc) + stmnt_runp (State m mi c assoc nbassoc f cycle pc) (Vblock lhs rhs) - (Runstate m mi c assoc' nbassoc f cycle pc) + (State m mi c assoc' nbassoc f cycle pc) | stmnt_runp_Vnonblock: forall lhs name rhs rhsval m mi c assoc nbassoc nbassoc' f cycle pc, assign_name lhs = OK name -> expr_runp (f cycle) assoc rhs rhsval -> nbassoc' = (PositiveMap.add name rhsval nbassoc) -> - stmnt_runp (Runstate m mi c assoc nbassoc f cycle pc) + stmnt_runp (State m mi c assoc nbassoc f cycle pc) (Vblock lhs rhs) - (Runstate m mi c assoc nbassoc' f cycle pc). + (State m mi c assoc nbassoc' f cycle pc). (*Fixpoint mi_step (s : state) (m : list module_item) {struct m} : res state := let run := fun st ls => @@ -509,8 +509,8 @@ Definition empty_assoclist : assoclist := PositiveMap.empty value. (*Definition mi_step_commit (s : state) (m : list module_item) : res state := match mi_step s m with - | OK (Runstate m assoc nbassoc f c) => - OK (Runstate m (merge_assoclist nbassoc assoc) empty_assoclist f c) + | OK (State m assoc nbassoc f c) => + OK (State m (merge_assoclist nbassoc assoc) empty_assoclist f c) | Error msg => Error msg | _ => Error (msg "Verilog: Wrong state") end.*) @@ -520,8 +520,8 @@ Definition empty_assoclist : assoclist := PositiveMap.empty value. match n with | S n' => do assoc' <- mi_run f assoc m n'; - match mi_step_commit (Runstate assoc' empty_assoclist f (Pos.of_nat n')) m with - | OK (Runstate assoc _ _ _) => OK assoc + match mi_step_commit (State assoc' empty_assoclist f (Pos.of_nat n')) m with + | OK (State assoc _ _ _) => OK assoc | Error m => Error m end | O => OK assoc @@ -583,22 +583,23 @@ Qed. Definition genv := Genv.t unit unit. + Inductive step : genv -> state -> Events.trace -> state -> Prop := | step_module : forall g m mi hmi tmi assoc nbassoc f cycle pc e, - step g (Runstate m mi (hmi::tmi) assoc nbassoc f cycle pc) e - (Runstate m hmi tmi assoc nbassoc f cycle pc) + step_cc g (State m mi (hmi::tmi) assoc nbassoc f cycle pc) e + (State m hmi tmi assoc nbassoc f cycle pc) | step_module_empty : forall g tmi hmi m mi assoc nbassoc f cycle pc e, hmi::tmi = mod_body m -> - step g (Runstate m mi nil assoc nbassoc f cycle pc) e - (Runstate m hmi tmi assoc nbassoc f (S cycle) pc). + step_cc g (State m mi nil assoc nbassoc f cycle pc) e + (State m hmi tmi assoc nbassoc f (S cycle) pc). Inductive initial_state (m: module): state -> Prop := | initial_state_intro: forall hmi tmi, hmi::tmi = mod_body m -> - initial_state m (Runstate m hmi tmi empty_assoclist empty_assoclist (initial_fextclk m) O xH). + initial_state m (State m hmi tmi empty_assoclist empty_assoclist (initial_fextclk m) O xH). (** A final state is a [Returnstate] with an empty call stack. *) -- cgit From 7ee0ca8263632536582646eb83b909f78f9e6fe4 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 3 May 2020 17:12:25 +0100 Subject: Add hex notation to values --- src/verilog/Value.v | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 241534c..1d31110 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -18,6 +18,7 @@ (* begin hide *) From bbv Require Import Word. +From bbv Require HexNotation WordScope. From Coq Require Import ZArith.ZArith. From compcert Require Import lib.Integers. (* end hide *) @@ -211,3 +212,11 @@ Definition shift_map (sz : nat) (f : word sz -> nat -> word sz) (w1 w2 : word sz Definition vshl v1 v2 := map_word2 (fun sz => shift_map sz (@wlshift sz)) v1 v2. Definition vshr v1 v2 := map_word2 (fun sz => shift_map sz (@wrshift sz)) v1 v2. + +Module HexNotationValue. + Export HexNotation. + Import WordScope. + + Notation "sz ''h' a" := (NToValue sz (hex a)) (at level 50). + +End HexNotationValue. -- cgit From 6a204dda8e46d2d491de599c9c0dc6bceed5e971 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 3 May 2020 17:12:39 +0100 Subject: Add state transition conversion functions --- src/translation/Veriloggen.v | 16 ++++++++++++++-- 1 file changed, 14 insertions(+), 2 deletions(-) diff --git a/src/translation/Veriloggen.v b/src/translation/Veriloggen.v index aa12a67..40876f7 100644 --- a/src/translation/Veriloggen.v +++ b/src/translation/Veriloggen.v @@ -38,6 +38,18 @@ Inductive statetrans : Type := | StateGoto (p : node) | StateCond (c : expr) (t f : node). +Definition state_goto (st : reg) (n : node) : stmnt := + Vnonblock (Vvar st) (Vlit (posToValue 32 n)). + +Definition state_cond (st : reg) (c : expr) (n1 n2 : node) : stmnt := + Vnonblock (Vvar st) (Vternary c (posToExpr 32 n1) (posToExpr 32 n2)). + +Definition statetrans_transl (stvar : reg) (st : statetrans) : stmnt := + match st with + | StateGoto p => state_goto stvar p + | StateCond c t f => state_cond stvar c t f + end. + Definition valid_freshstate (stm: PositiveMap.t stmnt) (fs: node) := forall (n: node), Plt n fs \/ stm!n = None. @@ -535,8 +547,8 @@ Definition make_stm (r : reg) (s : PositiveMap.t stmnt) : stmnt := Definition make_statetrans_cases (r : reg) (st : positive * statetrans) : expr * stmnt := match st with - | (n, StateGoto n') => (posToExpr 32 n, nonblock r (posToExpr 32 n')) - | (n, StateCond c n1 n2) => (posToExpr 32 n, nonblock r (Vternary c (posToExpr 32 n1) (posToExpr 32 n2))) + | (n, StateGoto n') => (posToExpr 32 n, state_goto r n') + | (n, StateCond c n1 n2) => (posToExpr 32 n, state_cond r c n1 n2) end. Definition make_statetrans (r : reg) (s : PositiveMap.t statetrans) : stmnt := -- cgit From d2c3d5a4fdb35b861be9df0795ef83f9b83c7bb7 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 3 May 2020 17:12:57 +0100 Subject: Add proofs and specification of Verilog conversion --- src/translation/Veriloggenproof.v | 46 ++++++++++++++++ src/translation/Veriloggenspec.v | 112 ++++++++++++++++++++++++++++++++++++++ 2 files changed, 158 insertions(+) create mode 100644 src/translation/Veriloggenproof.v create mode 100644 src/translation/Veriloggenspec.v diff --git a/src/translation/Veriloggenproof.v b/src/translation/Veriloggenproof.v new file mode 100644 index 0000000..cc884b4 --- /dev/null +++ b/src/translation/Veriloggenproof.v @@ -0,0 +1,46 @@ +(* -*- mode: coq -*- + * CoqUp: 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 . + *) + +From compcert Require Import Smallstep. +From compcert Require RTL. +From coqup Require Verilog. + +Section CORRECTNESS. + + Variable prog: RTL.program. + Variable tprog: Verilog.module. + + Inductive match_states: RTL.state -> Verilog.state -> Prop := + | match_state: + forall, + + match_states (RTL.State f s k sp e m) + (Verilog.State m mi mis assoc nbassoc f cycle pc) + | match_returnstate: + forall v tv k m tm cs + (MS: match_stacks k cs) + (LD: Val.lessdef v tv) + (MEXT: Mem.extends m tm), + match_states (CminorSel.Returnstate v k m) + (RTL.Returnstate cs tv tm). + + Theorem transf_program_correct: + forward_simulation (RTL.semantics prog) (Verilog.semantics tprog). + Admitted. + +End CORRECTNESS. diff --git a/src/translation/Veriloggenspec.v b/src/translation/Veriloggenspec.v new file mode 100644 index 0000000..67e2cad --- /dev/null +++ b/src/translation/Veriloggenspec.v @@ -0,0 +1,112 @@ +From Coq Require Import FSets.FMapPositive. +From compcert Require RTL Op Maps. +From coqup Require Import Coquplib Verilog Veriloggen Value. + +(** * Relational specification of the translation *) + +(** 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 *) + +Inductive tr_op : Op.operation -> list reg -> expr -> Prop := +| tr_op_Omove : forall r, tr_op Op.Omove (r::nil) (Vvar r) +| tr_op_Ointconst : forall n l, tr_op (Op.Ointconst n) l (Vlit (intToValue n)) +| tr_op_Oneg : forall r, tr_op Op.Oneg (r::nil) (Vunop Vneg (Vvar r)) +| tr_op_Osub : forall r1 r2, tr_op Op.Osub (r1::r2::nil) (bop Vsub r1 r2) +| tr_op_Omul : forall r1 r2, tr_op Op.Omul (r1::r2::nil) (bop Vmul r1 r2) +| tr_op_Omulimm : forall r n, tr_op (Op.Omulimm n) (r::nil) (boplit Vmul r n) +| tr_op_Odiv : forall r1 r2, tr_op Op.Odiv (r1::r2::nil) (bop Vdiv r1 r2) +| tr_op_Odivu : forall r1 r2, tr_op Op.Odivu (r1::r2::nil) (bop Vdivu r1 r2) +| tr_op_Omod : forall r1 r2, tr_op Op.Omod (r1::r2::nil) (bop Vmod r1 r2) +| tr_op_Omodu : forall r1 r2, tr_op Op.Omodu (r1::r2::nil) (bop Vmodu r1 r2) +| tr_op_Oand : forall r1 r2, tr_op Op.Oand (r1::r2::nil) (bop Vand r1 r2) +| tr_op_Oandimm : forall n r, tr_op (Op.Oandimm n) (r::nil) (boplit Vand r n) +| tr_op_Oor : forall r1 r2, tr_op Op.Oor (r1::r2::nil) (bop Vor r1 r2) +| tr_op_Oorimm : forall n r, tr_op (Op.Oorimm n) (r::nil) (boplit Vor r n) +| tr_op_Oxor : forall r1 r2, tr_op Op.Oxor (r1::r2::nil) (bop Vxor r1 r2) +| tr_op_Oxorimm : forall n r, tr_op (Op.Oxorimm n) (r::nil) (boplit Vxor r n) +| tr_op_Onot : forall r, tr_op Op.Onot (r::nil) (Vunop Vnot (Vvar r)) +| tr_op_Oshl : forall r1 r2, tr_op Op.Oshl (r1::r2::nil) (bop Vshl r1 r2) +| tr_op_Oshlimm : forall n r, tr_op (Op.Oshlimm n) (r::nil) (boplit Vshl r n) +| tr_op_Oshr : forall r1 r2, tr_op Op.Oshr (r1::r2::nil) (bop Vshr r1 r2) +| tr_op_Oshrimm : forall n r, tr_op (Op.Oshrimm n) (r::nil) (boplit Vshr r n) +| tr_op_Ocmp : forall c l e s s' i, translate_condition c l s = OK e s' i -> tr_op (Op.Ocmp c) l e +| tr_op_Olea : forall a l e s s' i, translate_eff_addressing a l s = OK e s' i -> tr_op (Op.Olea a) l e. + +Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Prop := +| tr_instr_Inop : + forall n, + tr_instr fin rtrn st (RTL.Inop n) Vskip (state_goto st n) +| tr_instr_Iop : + forall n op args e dst, + tr_op op args e -> + tr_instr fin rtrn st (RTL.Iop op args dst n) (Vnonblock (Vvar dst) e) (state_goto st n) +| tr_instr_Icond : + forall n1 n2 cond args s s' i c, + translate_condition cond args s = OK c s' i -> + tr_instr fin rtrn st (RTL.Icond cond args n1 n2) Vskip (state_cond st c n1 n2) +| tr_instr_Ireturn_None : + tr_instr fin rtrn st (RTL.Ireturn None) (block fin (Vlit (ZToValue 1%nat 1%Z))) Vskip +| tr_instr_Ireturn_Some : + forall r, + tr_instr fin rtrn st (RTL.Ireturn (Some r)) + (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r))) Vskip. + +Inductive tr_code (c : RTL.code) (stmnts trans : PositiveMap.t stmnt) + (fin rtrn st : reg) : Prop := +| tr_code_intro : + forall i s t n, + Maps.PTree.get n c = Some i -> + stmnts!n = Some s -> + trans!n = Some t -> + tr_instr fin rtrn st i s t -> + tr_code c stmnts trans fin rtrn st. + +(** [tr_module_items start clk st rst s mis] holds if the state is correctly +translated to actual Verilog, meaning it is correctly implemented as a state +machine in Verilog. Currently it does not seem that useful because it models +the [make_module_items] nearly exactly. Therefore it might have to be changed +to make up for that. *) + +Inductive tr_module_items : node -> reg -> reg -> reg -> state -> list module_item -> Prop := + tr_module_items_intro : + forall start clk st rst s mis, + Valways_ff (Vposedge clk) + (Vcond (Vbinop Veq (Vinputvar rst) (Vlit (ZToValue 1 1))) + (nonblock st (posToExpr 32 start)) + (make_statetrans st s.(st_statetrans))) + :: Valways_ff (Vposedge clk) (make_stm st s.(st_stm)) + :: List.map allocate_reg (PositiveMap.elements s.(st_decl)) = mis -> + tr_module_items start clk st rst s mis. + +(** [tr_module function module] Holds if the [module] is the correct translation +of the [function] that it was given. *) + +Inductive tr_module : RTL.function -> module -> Prop := + tr_module_intro : + forall f mi st rtrn fin clk rst start s0 s1 s2 s3 s4 s5 s6 i1 i2 i3 i4 i5 i6, + decl_io 1 s0 = OK fin s1 i1 -> + decl_io 32 s1 = OK rtrn s2 i2 -> + decl_io 1 s2 = OK start s3 i3 -> + decl_io 1 s3 = OK rst s4 i4 -> + decl_io 1 s4 = OK clk s5 i5 -> + decl_fresh_reg 32 s5 = OK st s6 i6 -> + tr_code f.(RTL.fn_code) s6.(st_stm) + (PositiveMap.map (statetrans_transl (fst st)) s6.(st_statetrans)) + (fst fin) (fst rtrn) (fst st) -> + tr_module_items f.(RTL.fn_entrypoint) (fst clk) (fst st) (fst rst) s6 mi -> + tr_module f (mkmodule + start + rst + clk + fin + rtrn + st + (List.map set_int_size f.(RTL.fn_params)) + mi). + +Lemma tr_module_correct: + forall f m, + transl_module f = Errors.OK m -> + tr_module f m. +Admitted. -- cgit From 0b118f8dca9068e0075e72f4d0c24cf707df44c7 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 3 May 2020 17:13:56 +0100 Subject: Add code to debug execution of HLS --- debug/CoqupTest.ml | 59 +++++++++++++++++++++++++++++++++++++++++++ debug/dune | 5 ++++ src/verilog/Test.v | 73 ++++++++++++++++++++++++++++++++++++++++++++++++++++++ 3 files changed, 137 insertions(+) create mode 100644 debug/CoqupTest.ml create mode 100644 debug/dune create mode 100644 src/verilog/Test.v diff --git a/debug/CoqupTest.ml b/debug/CoqupTest.ml new file mode 100644 index 0000000..f961f64 --- /dev/null +++ b/debug/CoqupTest.ml @@ -0,0 +1,59 @@ +open Coqup + +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 new file mode 100644 index 0000000..7f0c6a6 --- /dev/null +++ b/debug/dune @@ -0,0 +1,5 @@ +(include_subdirs no) + +(executable + (name CoqupTest) + (libraries coqup)) diff --git a/src/verilog/Test.v b/src/verilog/Test.v new file mode 100644 index 0000000..7a31865 --- /dev/null +++ b/src/verilog/Test.v @@ -0,0 +1,73 @@ +From coqup Require Import Verilog Veriloggen Coquplib Value. +From compcert Require Import AST Errors Maps Op Integers. +From compcert Require RTL. +From Coq Require Import FSets.FMapPositive. +Import ListNotations. +Import HexNotationValue. +Local Open Scope word_scope. + +Definition test_module : module := + let mods := [ + Valways (Vposedge 3%positive) (Vseq (Vnonblock (Vvar 6%positive) (Vlit (ZToValue 32 5))) + (Vnonblock (Vvar 7%positive) + (Vvar 6%positive))) + ] in + mkmodule (1%positive, 1%nat) + (2%positive, 1%nat) + (3%positive, 1%nat) + (4%positive, 1%nat) + (5%positive, 32%nat) + (6%positive, 32%nat) + nil + mods. + +Definition test_input : RTL.function := + let sig := mksignature nil Tvoid cc_default in + let params := nil in + let stacksize := 0 in + let entrypoint := 3%positive in + let code := PTree.set 1%positive (RTL.Ireturn (Some 1%positive)) + (PTree.set 3%positive (RTL.Iop (Ointconst (Int.repr 5)) nil 1%positive 1%positive) + (PTree.empty RTL.instruction)) in + RTL.mkfunction sig params stacksize code entrypoint. + +Definition test_input_program : RTL.program := + mkprogram [(1%positive, Gfun (Internal test_input))] nil 1%positive. + +Compute transf_program test_input_program. + +Definition test_output_module : module := + {| mod_start := (4%positive, 1%nat); + mod_reset := (5%positive, 1%nat); + mod_clk := (6%positive, 1%nat); + mod_finish := (2%positive, 1%nat); + mod_return := (3%positive, 32%nat); + mod_state := (7%positive, 32%nat); + mod_args := []; + mod_body := + [Valways_ff (Vposedge 6%positive) + (Vcond (Vbinop Veq (Vinputvar 5%positive) (1'h"1")) + (Vnonblock (Vvar 7%positive) (32'h"3")) + (Vcase (Vvar 7%positive) + [(Vlit (32'h"1"), Vnonblock (Vvar 7%positive) (32'h"1")); + (Vlit (32'h"3"), Vnonblock (Vvar 7%positive) (32'h"1"))] + (Some Vskip))); + Valways_ff (Vposedge 6%positive) + (Vcase (Vvar 7%positive) + [(Vlit (32'h"1"), Vseq (Vblock (Vvar 2%positive) (Vlit (1'h"1"))) + (Vblock (Vvar 3%positive) (Vvar 1%positive))); + (Vlit (32'h"3"), Vblock (Vvar 1%positive) (Vlit (32'h"5")))] + (Some Vskip)); + Vdecl 1%positive 32; Vdecl 7%positive 32] |}. + +Lemma valid_test_output : + transf_program test_input_program = OK test_output_module. +Proof. trivial. Qed. + +Lemma manual_simulation : + forall f, + step (State test_output_module empty_assoclist + (initial_fextclk test_output_module) + 1%nat 1%positive) + (State test_output_module test_output_module (add_assoclist 7 (32'h"3") empty_assoclist) + 2%nat 3%positive). -- cgit From c106b109fa0d469568c4841f07f7243a4f7813a4 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 4 May 2020 12:08:04 +0100 Subject: Refine the semantics --- src/verilog/Test.v | 37 +++++++++++++++--- src/verilog/Value.v | 44 +++++++++++++++++---- src/verilog/Verilog.v | 105 ++++++++++++++++++++++++++++++-------------------- 3 files changed, 130 insertions(+), 56 deletions(-) diff --git a/src/verilog/Test.v b/src/verilog/Test.v index 7a31865..5c22ef2 100644 --- a/src/verilog/Test.v +++ b/src/verilog/Test.v @@ -64,10 +64,35 @@ Lemma valid_test_output : transf_program test_input_program = OK test_output_module. Proof. trivial. Qed. +Definition test_fextclk := initial_fextclk test_output_module. + Lemma manual_simulation : - forall f, - step (State test_output_module empty_assoclist - (initial_fextclk test_output_module) - 1%nat 1%positive) - (State test_output_module test_output_module (add_assoclist 7 (32'h"3") empty_assoclist) - 2%nat 3%positive). + step (State test_output_module empty_assoclist empty_assoclist + test_fextclk 1 (32'h"1")) + (State test_output_module (add_assoclist 7%positive (32'h"3") empty_assoclist) + empty_assoclist test_fextclk 2 (32'h"3")). +Proof. + apply step_module with (assoc1 := empty_assoclist) + (nbassoc := (add_assoclist 7%positive (32'h"3") empty_assoclist)); auto. + apply mis_stepp_Cons with (s1 := State test_output_module empty_assoclist (add_assoclist 7%positive (32'h"3") empty_assoclist) test_fextclk 1%nat (32'h"1")). + apply mi_stepp_Valways_ff. + apply stmnt_runp_Vcond_true with (f := test_fextclk 1%nat) + (assoc := empty_assoclist) + (vc := 1'h"1"); auto. + apply get_state_fext_assoc. + apply erun_Vbinop with (lv := 1'h"1") + (rv := 1'h"1") + (oper := veq) + (EQ := EQ_trivial (1'h"1")); auto. + apply erun_Vinputvar; auto. + apply erun_Vlit. + eapply stmnt_runp_Vnonblock. simpl. auto. + apply erun_Vlit. + auto. + eapply mis_stepp_Cons. + apply mi_stepp_Valways_ff. + eapply stmnt_runp_Vcase_nomatch. + apply get_state_fext_assoc. + apply erun_Vvar_empty. auto. + apply erun_Vlit. + unfold ZToValue. instantiate (1 := 32%nat). simpl. diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 1d31110..21e59be 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -60,13 +60,6 @@ Definition valueToN (v : value) : N := Definition NToValue sz (n : N) : value := mkvalue sz (NToWord sz n). -Definition posToValue sz (p : positive) : value := - mkvalue sz (posToWord sz p). - -Definition posToValueAuto (p : positive) : value := - let size := Z.to_nat (Z.succ (log_inf p)) in - mkvalue size (Word.posToWord size p). - Definition ZToValue (s : nat) (z : Z) : value := mkvalue s (ZToWord s z). @@ -76,6 +69,19 @@ Definition valueToZ (v : value) : Z := Definition uvalueToZ (v : value) : Z := uwordToZ (vword v). +Definition posToValue sz (p : positive) : value := + mkvalue sz (posToWord sz p). + +Definition posToValueAuto (p : positive) : value := + let size := Z.to_nat (Z.succ (log_inf p)) in + mkvalue size (Word.posToWord size p). + +Definition valueToPos (v : value) : positive := + match valueToZ v with + | Zpos p => p + | _ => 1 + end. + Definition intToValue (i : Integers.int) : value := ZToValue Int.wordsize (Int.unsigned i). @@ -95,8 +101,10 @@ Definition boolToValue (sz : nat) (b : bool) : value := (** ** Arithmetic operations *) +Definition EQ_trivial : forall x, vsize x = vsize x. trivial. Defined. + Lemma unify_word (sz1 sz2 : nat) (w1 : word sz2): sz1 = sz2 -> word sz1. -Proof. intros; subst; assumption. Qed. +Proof. intros; subst; assumption. Defined. Definition value_eq_size: forall v1 v2 : value, { vsize v1 = vsize v2 } + { True }. @@ -141,6 +149,26 @@ Definition eq_to_opt (v1 v2 : value) (f : vsize v1 = vsize v2 -> value) | _ => None end. +Lemma eqvalue {sz : nat} (x y : word sz) : x = y -> (mkvalue sz x = mkvalue sz y). +Proof. intros. subst. reflexivity. Qed. + +Lemma nevalue {sz : nat} (x y : word sz) : x <> y -> (mkvalue sz x <> mkvalue sz y). +Proof. Admitted. + +Definition valueEqCheck (sz : nat) (x y : word sz) : + {mkvalue sz x = mkvalue sz y} + {mkvalue sz x <> mkvalue sz y} := + match weq x y with + | left eq => left (eqvalue x y eq) + | right ne => right (nevalue x y ne) + end. + +Definition valueEq (x y : value) (EQ : vsize x = vsize y): {x = y} + {x <> y} := + let unif_y := unify_word (vsize x) (vsize y) (vword y) EQ in + match weq (vword x) unif_y with + | left _ => left _ + | right _ => right _ + end. + (** Arithmetic operations over [value], interpreting them as signed or unsigned depending on the operation. diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 7a3982c..a927eac 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -31,6 +31,8 @@ From coqup Require Import common.Coquplib common.Show verilog.Value. From compcert Require Integers Events. From compcert Require Import Errors Smallstep Globalenvs. +Import HexNotationValue. + Notation "a ! b" := (PositiveMap.find b a) (at level 1). Definition reg : Type := positive. @@ -196,13 +198,11 @@ retrieved and set appropriately. Inductive state : Type := | State: forall (m : module) - (mi : module_item) - (mis : list module_item) (assoc : assoclist) (nbassoc : assoclist) (f : fextclk) (cycle : nat) - (pc : positive), + (stvar : value), state | Finishstate: forall v : value, @@ -248,21 +248,27 @@ Inductive expr_runp : fext -> assoclist -> expr -> value -> Prop := forall fext assoc v r, assoc!r = Some v -> expr_runp fext assoc (Vvar r) v + | erun_Vvar_empty : + forall fext assoc r sz, + assoc!r = None -> + expr_runp fext assoc (Vvar r) (ZToValue sz 0) | erun_Vinputvar : forall fext assoc r v, fext!r = Some v -> expr_runp fext assoc (Vinputvar r) v | erun_Vbinop : - forall fext assoc op l r lv rv oper EQ, + forall fext assoc op l r lv rv oper EQ resv, expr_runp fext assoc l lv -> expr_runp fext assoc r rv -> oper = binop_run op -> - expr_runp fext assoc (Vbinop op l r) (oper lv rv EQ) + resv = oper lv rv EQ -> + expr_runp fext assoc (Vbinop op l r) resv | erun_Vunop : - forall fext assoc u vu op oper, + forall fext assoc u vu op oper resv, expr_runp fext assoc u vu -> oper = unop_run op -> - expr_runp fext assoc (Vunop op u) (oper vu) + resv = oper vu -> + expr_runp fext assoc (Vunop op u) resv | erun_Vternary_true : forall fext assoc c ts fs vc vt, expr_runp fext assoc c vc -> @@ -402,17 +408,10 @@ Fixpoint stmnt_run' (n : nat) (s : state) (st : stmnt) {struct n} : res state := Definition stmnt_run (s : state) (st : stmnt) : res state := stmnt_run' (stmnt_height st) s st. *) -Inductive not_matched : state -> value -> expr * stmnt -> Prop := - not_in: - forall caseval curr currval m mi c assoc nbassoc f cycle pc, - expr_runp (f cycle) assoc (fst curr) currval -> - caseval <> currval -> - not_matched (State m mi c assoc nbassoc f cycle pc) caseval curr. - Inductive state_fext_assoc : state -> fext * assoclist -> Prop := | get_state_fext_assoc : - forall m mi c assoc nbassoc f cycle pc, - state_fext_assoc (State m mi c assoc nbassoc f cycle pc) (f cycle, assoc). + forall c assoc nbassoc f cycle pc, + state_fext_assoc (State c assoc nbassoc f cycle pc) (f cycle, assoc). Inductive stmnt_runp: state -> stmnt -> state -> Prop := | stmnt_run_Vskip: @@ -436,38 +435,44 @@ Inductive stmnt_runp: state -> stmnt -> state -> Prop := valueToBool vc = false -> stmnt_runp s0 stf s1 -> stmnt_runp s0 (Vcond c stt stf) s1 -| stmnt_runp_Vcase: - forall e ve s0 f assoc s1 me mve sc clst def, +| stmnt_runp_Vcase_match: + forall e ve s0 f assoc s1 me mve sc cs def, state_fext_assoc s0 (f, assoc) -> expr_runp f assoc e ve -> expr_runp f assoc me mve -> mve = ve -> - In (me, sc) clst -> stmnt_runp s0 sc s1 -> - stmnt_runp s0 (Vcase e clst def) s1 + stmnt_runp s0 (Vcase e ((me, sc)::cs) def) s1 +| stmnt_runp_Vcase_nomatch: + forall e ve s0 f assoc s1 me mve sc cs def, + state_fext_assoc s0 (f, assoc) -> + expr_runp f assoc e ve -> + expr_runp f assoc me mve -> + mve <> ve -> + stmnt_runp s0 (Vcase e cs def) s1 -> + stmnt_runp s0 (Vcase e ((me, sc)::cs) def) s1 | stmnt_runp_Vcase_default: - forall s0 f assoc st s1 e ve clst, + forall s0 f assoc st s1 e ve, state_fext_assoc s0 (f, assoc) -> expr_runp f assoc e ve -> - Forall (not_matched s0 ve) clst -> stmnt_runp s0 st s1 -> - stmnt_runp s0 (Vcase e clst (Some st)) s1 + stmnt_runp s0 (Vcase e nil (Some st)) s1 | stmnt_runp_Vblock: - forall lhs name rhs rhsval m mi c assoc assoc' nbassoc f cycle pc, + forall lhs name rhs rhsval c assoc assoc' nbassoc f cycle pc, assign_name lhs = OK name -> expr_runp (f cycle) assoc rhs rhsval -> assoc' = (PositiveMap.add name rhsval assoc) -> - stmnt_runp (State m mi c assoc nbassoc f cycle pc) + stmnt_runp (State c assoc nbassoc f cycle pc) (Vblock lhs rhs) - (State m mi c assoc' nbassoc f cycle pc) + (State c assoc' nbassoc f cycle pc) | stmnt_runp_Vnonblock: - forall lhs name rhs rhsval m mi c assoc nbassoc nbassoc' f cycle pc, + forall lhs name rhs rhsval c assoc nbassoc nbassoc' f cycle pc, assign_name lhs = OK name -> expr_runp (f cycle) assoc rhs rhsval -> nbassoc' = (PositiveMap.add name rhsval nbassoc) -> - stmnt_runp (State m mi c assoc nbassoc f cycle pc) - (Vblock lhs rhs) - (State m mi c assoc nbassoc' f cycle pc). + stmnt_runp (State c assoc nbassoc f cycle pc) + (Vnonblock lhs rhs) + (State c assoc nbassoc' f cycle pc). (*Fixpoint mi_step (s : state) (m : list module_item) {struct m} : res state := let run := fun st ls => @@ -499,6 +504,16 @@ Inductive mi_stepp : state -> module_item -> state -> Prop := forall s lhs rhs, mi_stepp s (Vdecl lhs rhs) s. +Inductive mis_stepp : state -> list module_item -> state -> Prop := +| mis_stepp_Cons : + forall mi mis s0 s1 s2, + mi_stepp s0 mi s1 -> + mis_stepp s1 mis s2 -> + mis_stepp s0 (mi :: mis) s2 +| mis_stepp_Nil : + forall s, + mis_stepp s nil s. + Definition add_assoclist (r : reg) (v : value) (assoc : assoclist) : assoclist := PositiveMap.add r v assoc. @@ -583,19 +598,24 @@ Qed. Definition genv := Genv.t unit unit. - -Inductive step : genv -> state -> Events.trace -> state -> Prop := +Inductive step : state -> state -> Prop := | step_module : - forall g m mi hmi tmi assoc nbassoc f cycle pc e, - step_cc g (State m mi (hmi::tmi) assoc nbassoc f cycle pc) e - (State m hmi tmi assoc nbassoc f cycle pc) -| step_module_empty : - forall g tmi hmi m mi assoc nbassoc f cycle pc e, - hmi::tmi = mod_body m -> - step_cc g (State m mi nil assoc nbassoc f cycle pc) e - (State m hmi tmi assoc nbassoc f (S cycle) pc). - -Inductive initial_state (m: module): state -> Prop := + forall m stvar stvar' cycle f assoc0 assoc1 assoc2 nbassoc, + mis_stepp (State m assoc0 empty_assoclist f cycle stvar) + m.(mod_body) + (State m assoc1 nbassoc f cycle stvar) -> + assoc2 = merge_assoclist nbassoc assoc1 -> + Some stvar' = assoc2!(fst m.(mod_state)) -> + step (State m assoc0 empty_assoclist f cycle stvar) + (State m assoc2 empty_assoclist f (S cycle) stvar') +| step_finish : + forall m assoc f cycle stvar result, + assoc!(fst m.(mod_finish)) = Some (1'h"1") -> + assoc!(fst m.(mod_return)) = Some result -> + step (State m assoc empty_assoclist f cycle stvar) + (Finishstate result). + +(*Inductive initial_state (m: module): state -> Prop := | initial_state_intro: forall hmi tmi, hmi::tmi = mod_body m -> @@ -611,3 +631,4 @@ Inductive final_state: state -> Integers.int -> Prop := Definition semantics (p: module) := Semantics step (initial_state p) final_state (Genv.empty_genv unit unit nil). +*) -- cgit From dfef69f7844a67210032bfe2597ba5d6b6702469 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 4 May 2020 16:08:43 +0100 Subject: Move Verilog to .sv --- example/main.sv | 77 +++++++++++++++++++++++++++++++++++++++++++++++++++++++ example/main.v | 79 --------------------------------------------------------- 2 files changed, 77 insertions(+), 79 deletions(-) create mode 100644 example/main.sv delete mode 100644 example/main.v diff --git a/example/main.sv b/example/main.sv new file mode 100644 index 0000000..8e97a43 --- /dev/null +++ b/example/main.sv @@ -0,0 +1,77 @@ +module main(input start, reset, clk, + output finished, output [31:0] return_val); + + reg [31:0] x; + reg [31:0] y; + reg [31:0] z; + reg [2:0] state; + + reg [31:0] return_val_w; + reg finished_w; + + localparam [2:0] START_STATE = 0; + localparam [2:0] MAIN_STATE_0 = 1; + localparam [2:0] MAIN_STATE_1 = 2; + localparam [2:0] MAIN_STATE_2 = 3; + localparam [2:0] MAIN_STATE_3 = 4; + localparam [2:0] MAIN_STATE_4 = 5; + localparam [2:0] FINISHED_STATE = 6; + + assign return_val = return_val_w; + assign finished = finished_w; + + always @(posedge clk) + if (reset) begin + state <= START_STATE; + return_val_w <= 0; + finished_w <= 0; + end + else + case (state) + START_STATE: x <= 0; + MAIN_STATE_0: y <= 0; + MAIN_STATE_1: z <= 0; + MAIN_STATE_2: y <= 2; + MAIN_STATE_3: z <= 3; + MAIN_STATE_4: x <= y * z; + FINISHED_STATE: begin + return_val_w <= x; + finished_w <= 1; + end + default: state <= START_STATE; + endcase + + always @(posedge clk) + if (state != FINISHED_STATE) + state <= state + 1; + +endmodule + +module testbench; + reg start, reset, clk; + wire finished; + wire [31:0] return_val; + + main main(start, reset, clk, finished, return_val); + + initial begin + $dumpvars; + start = 0; + reset = 1; + clk = 0; + + @(posedge clk) begin + reset = 0; + start = 1; + end + @(posedge clk) start = 0; + + #100; + + $display("Result: %d", return_val); + $finish; + end + + always #5 clk = ~clk; + +endmodule diff --git a/example/main.v b/example/main.v deleted file mode 100644 index 63130fe..0000000 --- a/example/main.v +++ /dev/null @@ -1,79 +0,0 @@ -// -*- mode: verilog -*- - -module main(input start, reset, clk, - output finished, output [31:0] return_val); - - reg [31:0] x; - reg [31:0] y; - reg [31:0] z; - reg [2:0] state; - - reg [31:0] return_val_w; - reg finished_w; - - localparam [2:0] START_STATE = 0; - localparam [2:0] MAIN_STATE_0 = 1; - localparam [2:0] MAIN_STATE_1 = 2; - localparam [2:0] MAIN_STATE_2 = 3; - localparam [2:0] MAIN_STATE_3 = 4; - localparam [2:0] MAIN_STATE_4 = 5; - localparam [2:0] FINISHED_STATE = 6; - - assign return_val = return_val_w; - assign finished = finished_w; - - always @(posedge clk) - if (reset) begin - state <= START_STATE; - return_val_w <= 0; - finished_w <= 0; - end - else - case (state) - START_STATE: x <= 0; - MAIN_STATE_0: y <= 0; - MAIN_STATE_1: z <= 0; - MAIN_STATE_2: y <= 2; - MAIN_STATE_3: z <= 3; - MAIN_STATE_4: x <= y * z; - FINISHED_STATE: begin - return_val_w <= x; - finished_w <= 1; - end - default: state <= START_STATE; - endcase - - always @(posedge clk) - if (state != FINISHED_STATE) - state <= state + 1; - -endmodule - -module testbench; - reg start, reset, clk; - wire finished; - wire [31:0] return_val; - - main main(start, reset, clk, finished, return_val); - - initial begin - $dumpvars; - start = 0; - reset = 1; - clk = 0; - - @(posedge clk) begin - reset = 0; - start = 1; - end - @(posedge clk) start = 0; - - #100; - - $display("Result: %d", return_val); - $finish; - end - - always #5 clk = ~clk; - -endmodule -- cgit From 5cfa4af9c6e4d9703e3142c24ae78c7da0ac575f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 4 May 2020 16:13:16 +0100 Subject: Add equality check for value --- src/Compiler.v | 2 +- src/Simulator.v | 2 +- src/extraction/Extraction.v | 2 +- src/translation/Veriloggen.v | 2 +- src/translation/Veriloggenproof.v | 2 +- src/verilog/Value.v | 36 +++++++++++++++++++++--------------- src/verilog/Verilog.v | 2 +- 7 files changed, 27 insertions(+), 21 deletions(-) diff --git a/src/Compiler.v b/src/Compiler.v index 697732d..e998521 100644 --- a/src/Compiler.v +++ b/src/Compiler.v @@ -1,4 +1,4 @@ -(* -*- mode: coq -*- +(* * CoqUp: Verified high-level synthesis. * Copyright (C) 2019-2020 Yann Herklotz * diff --git a/src/Simulator.v b/src/Simulator.v index 6849f80..930971b 100644 --- a/src/Simulator.v +++ b/src/Simulator.v @@ -1,4 +1,4 @@ -(* -*- mode: coq -*- +(* * CoqUp: Verified high-level synthesis. * Copyright (C) 2020 Yann Herklotz * diff --git a/src/extraction/Extraction.v b/src/extraction/Extraction.v index 319e2eb..0028fcb 100644 --- a/src/extraction/Extraction.v +++ b/src/extraction/Extraction.v @@ -1,4 +1,4 @@ -(* -*- mode: coq -*- +(* * CoqUp: Verified high-level synthesis. * Copyright (C) 2019-2020 Yann Herklotz * diff --git a/src/translation/Veriloggen.v b/src/translation/Veriloggen.v index 40876f7..f127b11 100644 --- a/src/translation/Veriloggen.v +++ b/src/translation/Veriloggen.v @@ -1,4 +1,4 @@ -(* -*- mode: coq -*- +(* * CoqUp: Verified high-level synthesis. * Copyright (C) 2020 Yann Herklotz * diff --git a/src/translation/Veriloggenproof.v b/src/translation/Veriloggenproof.v index cc884b4..942a83a 100644 --- a/src/translation/Veriloggenproof.v +++ b/src/translation/Veriloggenproof.v @@ -1,4 +1,4 @@ -(* -*- mode: coq -*- +(* * CoqUp: Verified high-level synthesis. * Copyright (C) 2020 Yann Herklotz * diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 21e59be..4cacab5 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -1,4 +1,4 @@ -(* -*- mode: coq -*- +(* * CoqUp: Verified high-level synthesis. * Copyright (C) 2020 Yann Herklotz * @@ -149,24 +149,30 @@ Definition eq_to_opt (v1 v2 : value) (f : vsize v1 = vsize v2 -> value) | _ => None end. -Lemma eqvalue {sz : nat} (x y : word sz) : x = y -> (mkvalue sz x = mkvalue sz y). -Proof. intros. subst. reflexivity. Qed. -Lemma nevalue {sz : nat} (x y : word sz) : x <> y -> (mkvalue sz x <> mkvalue sz y). -Proof. Admitted. +Lemma eqvalue {sz : nat} (x y : word sz) : x = y <-> mkvalue sz x = mkvalue sz y. +Proof. + split; intros. + subst. reflexivity. inversion H. apply existT_wordToZ in H1. + apply wordToZ_inj. assumption. +Qed. + +Lemma eqvaluef {sz : nat} (x y : word sz) : x = y -> mkvalue sz x = mkvalue sz y. +Proof. apply eqvalue. Qed. + +Lemma nevalue {sz : nat} (x y : word sz) : x <> y <-> mkvalue sz x <> mkvalue sz y. +Proof. split; intros; intuition. apply H. apply eqvalue. assumption. + apply H. rewrite H0. trivial. +Qed. -Definition valueEqCheck (sz : nat) (x y : word sz) : +Lemma nevaluef {sz : nat} (x y : word sz) : x <> y -> mkvalue sz x <> mkvalue sz y. +Proof. apply nevalue. Qed. + +Definition valueEq (sz : nat) (x y : word sz) : {mkvalue sz x = mkvalue sz y} + {mkvalue sz x <> mkvalue sz y} := match weq x y with - | left eq => left (eqvalue x y eq) - | right ne => right (nevalue x y ne) - end. - -Definition valueEq (x y : value) (EQ : vsize x = vsize y): {x = y} + {x <> y} := - let unif_y := unify_word (vsize x) (vsize y) (vword y) EQ in - match weq (vword x) unif_y with - | left _ => left _ - | right _ => right _ + | left eq => left (eqvaluef x y eq) + | right ne => right (nevaluef x y ne) end. (** Arithmetic operations over [value], interpreting them as signed or unsigned diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index a927eac..756dc12 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -1,4 +1,4 @@ -(* -*- mode: coq -*- +(* * CoqUp: Verified high-level synthesis. * Copyright (C) 2019-2020 Yann Herklotz * -- cgit From 2df5dc835efa3ac7552363e81ef9af5d3145cc7e Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 5 May 2020 00:19:27 +0100 Subject: Finish manual simulation --- src/verilog/Test.v | 44 +++++++++++++++++++++++++++++++++++++++++++- src/verilog/Value.v | 29 +++++++++++++++++++++++++---- 2 files changed, 68 insertions(+), 5 deletions(-) diff --git a/src/verilog/Test.v b/src/verilog/Test.v index 5c22ef2..5fd6d07 100644 --- a/src/verilog/Test.v +++ b/src/verilog/Test.v @@ -1,9 +1,30 @@ +(* + * CoqUp: 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 . + *) + From coqup Require Import Verilog Veriloggen Coquplib Value. From compcert Require Import AST Errors Maps Op Integers. From compcert Require RTL. From Coq Require Import FSets.FMapPositive. +From bbv Require Import Word. Import ListNotations. Import HexNotationValue. +Import WordScope. + Local Open Scope word_scope. Definition test_module : module := @@ -60,6 +81,8 @@ Definition test_output_module : module := (Some Vskip)); Vdecl 1%positive 32; Vdecl 7%positive 32] |}. +Search (_ <> _ -> _ = _). + Lemma valid_test_output : transf_program test_input_program = OK test_output_module. Proof. trivial. Qed. @@ -95,4 +118,23 @@ Proof. apply get_state_fext_assoc. apply erun_Vvar_empty. auto. apply erun_Vlit. - unfold ZToValue. instantiate (1 := 32%nat). simpl. + unfold ZToValue. instantiate (1 := 32%nat). simpl. apply nevalue. simpl. + apply weqb_false. simpl. trivial. + eapply stmnt_runp_Vcase_nomatch. + apply get_state_fext_assoc. + apply erun_Vvar_empty. auto. + apply erun_Vlit. + unfold ZToValue. instantiate (1 := 32%nat). simpl. apply nevalue. simpl. + apply weqb_false. simpl. trivial. + eapply stmnt_runp_Vcase_default. + apply get_state_fext_assoc. + apply erun_Vvar_empty. auto. + apply stmnt_run_Vskip. + eapply mis_stepp_Cons. + apply mi_stepp_Vdecl. + eapply mis_stepp_Cons. + apply mi_stepp_Vdecl. + apply mis_stepp_Nil. + Unshelve. + exact 0%nat. +Qed. diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 4cacab5..39d1832 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -103,8 +103,8 @@ Definition boolToValue (sz : nat) (b : bool) : value := Definition EQ_trivial : forall x, vsize x = vsize x. trivial. Defined. -Lemma unify_word (sz1 sz2 : nat) (w1 : word sz2): sz1 = sz2 -> word sz1. -Proof. intros; subst; assumption. Defined. +Definition unify_word (sz1 sz2 : nat) (w1 : word sz2): sz1 = sz2 -> word sz1. +intros; subst; assumption. Defined. Definition value_eq_size: forall v1 v2 : value, { vsize v1 = vsize v2 } + { True }. @@ -149,7 +149,6 @@ Definition eq_to_opt (v1 v2 : value) (f : vsize v1 = vsize v2 -> value) | _ => None end. - Lemma eqvalue {sz : nat} (x y : word sz) : x = y <-> mkvalue sz x = mkvalue sz y. Proof. split; intros. @@ -168,13 +167,35 @@ Qed. Lemma nevaluef {sz : nat} (x y : word sz) : x <> y -> mkvalue sz x <> mkvalue sz y. Proof. apply nevalue. Qed. -Definition valueEq (sz : nat) (x y : word sz) : +(*Definition rewrite_word_size (initsz finalsz : nat) (w : word initsz) + : option (word finalsz) := + match Nat.eqb initsz finalsz return option (word finalsz) with + | true => Some _ + | false => None + end.*) + +Definition valueeq (sz : nat) (x y : word sz) : {mkvalue sz x = mkvalue sz y} + {mkvalue sz x <> mkvalue sz y} := match weq x y with | left eq => left (eqvaluef x y eq) | right ne => right (nevaluef x y ne) end. +Definition valueeqb (x y : value) : bool := + match value_eq_size x y with + | left EQ => + weqb (vword x) (unify_word (vsize x) (vsize y) (vword y) EQ) + | right _ => false + end. + +Theorem valueeqb_true_iff : + forall v1 v2, + valueeqb v1 v2 = true <-> v1 = v2. +Proof. + split; intros. + unfold valueeqb in H. destruct (value_eq_size v1 v2). + Admitted. + (** Arithmetic operations over [value], interpreting them as signed or unsigned depending on the operation. -- cgit From 5d4aed9030565349a83f9ede4bd35c020eb416b5 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 5 May 2020 00:38:19 +0100 Subject: Simplifications to proof --- src/verilog/Test.v | 21 +++++++-------------- src/verilog/Value.v | 2 -- src/verilog/Verilog.v | 10 ++++++++-- 3 files changed, 15 insertions(+), 18 deletions(-) diff --git a/src/verilog/Test.v b/src/verilog/Test.v index 5fd6d07..9e13bf6 100644 --- a/src/verilog/Test.v +++ b/src/verilog/Test.v @@ -95,20 +95,14 @@ Lemma manual_simulation : (State test_output_module (add_assoclist 7%positive (32'h"3") empty_assoclist) empty_assoclist test_fextclk 2 (32'h"3")). Proof. - apply step_module with (assoc1 := empty_assoclist) - (nbassoc := (add_assoclist 7%positive (32'h"3") empty_assoclist)); auto. - apply mis_stepp_Cons with (s1 := State test_output_module empty_assoclist (add_assoclist 7%positive (32'h"3") empty_assoclist) test_fextclk 1%nat (32'h"1")). + eapply step_module; auto. + eapply mis_stepp_Cons; auto. apply mi_stepp_Valways_ff. - apply stmnt_runp_Vcond_true with (f := test_fextclk 1%nat) - (assoc := empty_assoclist) - (vc := 1'h"1"); auto. + eapply stmnt_runp_Vcond_true; auto. apply get_state_fext_assoc. - apply erun_Vbinop with (lv := 1'h"1") - (rv := 1'h"1") - (oper := veq) - (EQ := EQ_trivial (1'h"1")); auto. - apply erun_Vinputvar; auto. - apply erun_Vlit. + eapply erun_Vbinop; auto. + apply erun_Vinputvar; simpl. auto. + apply erun_Vlit. simpl. unfold ZToValue. simpl. instantiate (1 := eq_refl (vsize (1'h"1"))). auto. eapply stmnt_runp_Vnonblock. simpl. auto. apply erun_Vlit. auto. @@ -135,6 +129,5 @@ Proof. eapply mis_stepp_Cons. apply mi_stepp_Vdecl. apply mis_stepp_Nil. - Unshelve. - exact 0%nat. + auto. Unshelve. exact 0%nat. Qed. diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 39d1832..61f2652 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -101,8 +101,6 @@ Definition boolToValue (sz : nat) (b : bool) : value := (** ** Arithmetic operations *) -Definition EQ_trivial : forall x, vsize x = vsize x. trivial. Defined. - Definition unify_word (sz1 sz2 : nat) (w1 : word sz2): sz1 = sz2 -> word sz1. intros; subst; assumption. Defined. diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 756dc12..399bff7 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -281,6 +281,7 @@ Inductive expr_runp : fext -> assoclist -> expr -> value -> Prop := expr_runp fext assoc fs vf -> valueToBool vc = false -> expr_runp fext assoc (Vternary c ts fs) vf. +Hint Constructors expr_runp : verilog. Definition handle_opt {A : Type} (err : errmsg) (val : option A) : res A := @@ -412,11 +413,12 @@ Inductive state_fext_assoc : state -> fext * assoclist -> Prop := | get_state_fext_assoc : forall c assoc nbassoc f cycle pc, state_fext_assoc (State c assoc nbassoc f cycle pc) (f cycle, assoc). +Hint Constructors state_fext_assoc : verilog. Inductive stmnt_runp: state -> stmnt -> state -> Prop := -| stmnt_run_Vskip: +| stmnt_runp_Vskip: forall s, stmnt_runp s Vskip s -| stmnt_run_Vseq: +| stmnt_runp_Vseq: forall st1 st2 s0 s1 s2, stmnt_runp s0 st1 s1 -> stmnt_runp s1 st2 s2 -> @@ -473,6 +475,7 @@ Inductive stmnt_runp: state -> stmnt -> state -> Prop := stmnt_runp (State c assoc nbassoc f cycle pc) (Vnonblock lhs rhs) (State c assoc nbassoc' f cycle pc). +Hint Constructors stmnt_runp : verilog. (*Fixpoint mi_step (s : state) (m : list module_item) {struct m} : res state := let run := fun st ls => @@ -503,6 +506,7 @@ Inductive mi_stepp : state -> module_item -> state -> Prop := | mi_stepp_Vdecl : forall s lhs rhs, mi_stepp s (Vdecl lhs rhs) s. +Hint Constructors mi_stepp : verilog. Inductive mis_stepp : state -> list module_item -> state -> Prop := | mis_stepp_Cons : @@ -513,6 +517,7 @@ Inductive mis_stepp : state -> list module_item -> state -> Prop := | mis_stepp_Nil : forall s, mis_stepp s nil s. +Hint Constructors mis_stepp : verilog. Definition add_assoclist (r : reg) (v : value) (assoc : assoclist) : assoclist := PositiveMap.add r v assoc. @@ -614,6 +619,7 @@ Inductive step : state -> state -> Prop := assoc!(fst m.(mod_return)) = Some result -> step (State m assoc empty_assoclist f cycle stvar) (Finishstate result). +Hint Constructors step : verilog. (*Inductive initial_state (m: module): state -> Prop := | initial_state_intro: -- cgit From 4d409fe68c693dd083295f35bd5af8bdec6d2632 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 5 May 2020 01:16:03 +0100 Subject: Minimised manual simulation --- src/verilog/Test.v | 39 ++++----------------------------------- src/verilog/Verilog.v | 20 ++++++++++---------- 2 files changed, 14 insertions(+), 45 deletions(-) diff --git a/src/verilog/Test.v b/src/verilog/Test.v index 9e13bf6..d94467d 100644 --- a/src/verilog/Test.v +++ b/src/verilog/Test.v @@ -95,39 +95,8 @@ Lemma manual_simulation : (State test_output_module (add_assoclist 7%positive (32'h"3") empty_assoclist) empty_assoclist test_fextclk 2 (32'h"3")). Proof. - eapply step_module; auto. - eapply mis_stepp_Cons; auto. - apply mi_stepp_Valways_ff. - eapply stmnt_runp_Vcond_true; auto. - apply get_state_fext_assoc. - eapply erun_Vbinop; auto. - apply erun_Vinputvar; simpl. auto. - apply erun_Vlit. simpl. unfold ZToValue. simpl. instantiate (1 := eq_refl (vsize (1'h"1"))). auto. - eapply stmnt_runp_Vnonblock. simpl. auto. - apply erun_Vlit. - auto. - eapply mis_stepp_Cons. - apply mi_stepp_Valways_ff. - eapply stmnt_runp_Vcase_nomatch. - apply get_state_fext_assoc. - apply erun_Vvar_empty. auto. - apply erun_Vlit. - unfold ZToValue. instantiate (1 := 32%nat). simpl. apply nevalue. simpl. - apply weqb_false. simpl. trivial. - eapply stmnt_runp_Vcase_nomatch. - apply get_state_fext_assoc. - apply erun_Vvar_empty. auto. - apply erun_Vlit. - unfold ZToValue. instantiate (1 := 32%nat). simpl. apply nevalue. simpl. - apply weqb_false. simpl. trivial. - eapply stmnt_runp_Vcase_default. - apply get_state_fext_assoc. - apply erun_Vvar_empty. auto. - apply stmnt_run_Vskip. - eapply mis_stepp_Cons. - apply mi_stepp_Vdecl. - eapply mis_stepp_Cons. - apply mi_stepp_Vdecl. - apply mis_stepp_Nil. - auto. Unshelve. exact 0%nat. + repeat (econstructor; eauto); + try (simpl; unfold ZToValue; simpl; instantiate (1 := eq_refl (vsize (1'h"1"))); auto); + apply nevalue; apply weqb_false; trivial. + Unshelve. exact 0%nat. Qed. diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 399bff7..451b452 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -244,14 +244,14 @@ Inductive expr_runp : fext -> assoclist -> expr -> value -> Prop := | erun_Vlit : forall fext assoc v, expr_runp fext assoc (Vlit v) v - | erun_Vvar : - forall fext assoc v r, - assoc!r = Some v -> - expr_runp fext assoc (Vvar r) v | erun_Vvar_empty : forall fext assoc r sz, assoc!r = None -> expr_runp fext assoc (Vvar r) (ZToValue sz 0) + | erun_Vvar : + forall fext assoc v r, + assoc!r = Some v -> + expr_runp fext assoc (Vvar r) v | erun_Vinputvar : forall fext assoc r v, fext!r = Some v -> @@ -437,21 +437,21 @@ Inductive stmnt_runp: state -> stmnt -> state -> Prop := valueToBool vc = false -> stmnt_runp s0 stf s1 -> stmnt_runp s0 (Vcond c stt stf) s1 -| stmnt_runp_Vcase_match: +| stmnt_runp_Vcase_nomatch: forall e ve s0 f assoc s1 me mve sc cs def, state_fext_assoc s0 (f, assoc) -> expr_runp f assoc e ve -> expr_runp f assoc me mve -> - mve = ve -> - stmnt_runp s0 sc s1 -> + mve <> ve -> + stmnt_runp s0 (Vcase e cs def) s1 -> stmnt_runp s0 (Vcase e ((me, sc)::cs) def) s1 -| stmnt_runp_Vcase_nomatch: +| stmnt_runp_Vcase_match: forall e ve s0 f assoc s1 me mve sc cs def, state_fext_assoc s0 (f, assoc) -> expr_runp f assoc e ve -> expr_runp f assoc me mve -> - mve <> ve -> - stmnt_runp s0 (Vcase e cs def) s1 -> + mve = ve -> + stmnt_runp s0 sc s1 -> stmnt_runp s0 (Vcase e ((me, sc)::cs) def) s1 | stmnt_runp_Vcase_default: forall s0 f assoc st s1 e ve, -- cgit From d530dc044242efdd4c5fafe13456cd05ac65beee Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 5 May 2020 01:20:58 +0100 Subject: Refine test file --- src/verilog/Test.v | 7 ++----- 1 file changed, 2 insertions(+), 5 deletions(-) diff --git a/src/verilog/Test.v b/src/verilog/Test.v index d94467d..5cd6ec4 100644 --- a/src/verilog/Test.v +++ b/src/verilog/Test.v @@ -81,8 +81,6 @@ Definition test_output_module : module := (Some Vskip)); Vdecl 1%positive 32; Vdecl 7%positive 32] |}. -Search (_ <> _ -> _ = _). - Lemma valid_test_output : transf_program test_input_program = OK test_output_module. Proof. trivial. Qed. @@ -96,7 +94,6 @@ Lemma manual_simulation : empty_assoclist test_fextclk 2 (32'h"3")). Proof. repeat (econstructor; eauto); - try (simpl; unfold ZToValue; simpl; instantiate (1 := eq_refl (vsize (1'h"1"))); auto); - apply nevalue; apply weqb_false; trivial. - Unshelve. exact 0%nat. + try (unfold ZToValue; instantiate (1 := eq_refl (vsize (1'h"1"))); auto); + apply nevalue; apply weqb_false; trivial. Unshelve. exact 0%nat. Qed. -- cgit From 74ac24f8dc099cc558d3b03b2f9303c89048f519 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 6 May 2020 11:49:32 +0100 Subject: Add changes to value --- src/verilog/Value.v | 11 +++++++++-- 1 file changed, 9 insertions(+), 2 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 61f2652..8a0716d 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -186,13 +186,20 @@ Definition valueeqb (x y : value) : bool := | right _ => false end. +Definition value_projZ_eqb (v1 v2 : value) : bool := Z.eqb (valueToZ v1) (valueToZ v2). + +Theorem value_projZ_eqb_true : + forall v1 v2, + v1 = v2 -> value_projZ_eqb v1 v2 = true. +Proof. intros. subst. unfold value_projZ_eqb. apply Z.eqb_eq. trivial. Qed. + Theorem valueeqb_true_iff : forall v1 v2, valueeqb v1 v2 = true <-> v1 = v2. Proof. split; intros. - unfold valueeqb in H. destruct (value_eq_size v1 v2). - Admitted. + unfold valueeqb in H. destruct (value_eq_size v1 v2) eqn:?. + - destruct v1, v2. simpl in H. (** Arithmetic operations over [value], interpreting them as signed or unsigned depending on the operation. -- cgit From 1d6589fbda24fb57c9ab74d4bd41fa1a487a2c62 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 7 May 2020 17:21:37 +0100 Subject: Remove Admitted Maps Lemma --- src/common/Maps.v | 6 ------ 1 file changed, 6 deletions(-) diff --git a/src/common/Maps.v b/src/common/Maps.v index 7a542b7..3236417 100644 --- a/src/common/Maps.v +++ b/src/common/Maps.v @@ -40,10 +40,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). -Lemma traverse1_inversion: - forall (A B : Type) (f : A -> res B) (i j : positive) (m : t A) (m' : t B), - traverse1 f m = OK m' -> - list_forall2 (fun x y => f (snd x) = OK (snd y) /\ fst x = fst y) (elements m) (elements m'). -Proof. Admitted. - End PTree. -- cgit From c8535e1f454c7014c7b794fc8be343e2fda97937 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 7 May 2020 17:21:51 +0100 Subject: Rename assoclist to assocset --- src/verilog/Test.v | 6 +++--- src/verilog/Verilog.v | 50 +++++++++++++++++++++++++------------------------- 2 files changed, 28 insertions(+), 28 deletions(-) diff --git a/src/verilog/Test.v b/src/verilog/Test.v index 5cd6ec4..f24612e 100644 --- a/src/verilog/Test.v +++ b/src/verilog/Test.v @@ -88,10 +88,10 @@ Proof. trivial. Qed. Definition test_fextclk := initial_fextclk test_output_module. Lemma manual_simulation : - step (State test_output_module empty_assoclist empty_assoclist + step (State test_output_module empty_assocset empty_assocset test_fextclk 1 (32'h"1")) - (State test_output_module (add_assoclist 7%positive (32'h"3") empty_assoclist) - empty_assoclist test_fextclk 2 (32'h"3")). + (State test_output_module (add_assocset 7%positive (32'h"3") empty_assocset) + empty_assocset test_fextclk 2 (32'h"3")). Proof. repeat (econstructor; eauto); try (unfold ZToValue; instantiate (1 := eq_refl (vsize (1'h"1"))); auto); diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index 451b452..a35fb58 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -39,7 +39,7 @@ Definition reg : Type := positive. Definition node : Type := positive. Definition szreg : Type := reg * nat. -Definition assoclist : Type := PositiveMap.t value. +Definition assocset : Type := PositiveMap.t value. (** * Verilog AST @@ -52,7 +52,7 @@ 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. *) -Definition estate : Type := assoclist * assoclist. +Definition estate : Type := assocset * assocset. (** ** Binary Operators @@ -198,8 +198,8 @@ retrieved and set appropriately. Inductive state : Type := | State: forall (m : module) - (assoc : assoclist) - (nbassoc : assoclist) + (assoc : assocset) + (nbassoc : assocset) (f : fextclk) (cycle : nat) (stvar : value), @@ -240,7 +240,7 @@ Definition unop_run (op : unop) : value -> value := | Vnot => vbitneg end. -Inductive expr_runp : fext -> assoclist -> expr -> value -> Prop := +Inductive expr_runp : fext -> assocset -> expr -> value -> Prop := | erun_Vlit : forall fext assoc v, expr_runp fext assoc (Vlit v) v @@ -306,7 +306,7 @@ Definition access_fext (f : fext) (r : reg) : res value := end. (* TODO FIX Vvar case without default *) -(*Fixpoint expr_run (assoc : assoclist) (e : expr) +(*Fixpoint expr_run (assoc : assocset) (e : expr) {struct e} : res value := match e with | Vlit v => OK v @@ -409,7 +409,7 @@ Fixpoint stmnt_run' (n : nat) (s : state) (st : stmnt) {struct n} : res state := Definition stmnt_run (s : state) (st : stmnt) : res state := stmnt_run' (stmnt_height st) s st. *) -Inductive state_fext_assoc : state -> fext * assoclist -> Prop := +Inductive state_fext_assoc : state -> fext * assocset -> Prop := | get_state_fext_assoc : forall c assoc nbassoc f cycle pc, state_fext_assoc (State c assoc nbassoc f cycle pc) (f cycle, assoc). @@ -519,28 +519,28 @@ Inductive mis_stepp : state -> list module_item -> state -> Prop := mis_stepp s nil s. Hint Constructors mis_stepp : verilog. -Definition add_assoclist (r : reg) (v : value) (assoc : assoclist) : assoclist := +Definition add_assocset (r : reg) (v : value) (assoc : assocset) : assocset := PositiveMap.add r v assoc. -Definition merge_assoclist (nbassoc assoc : assoclist) : assoclist := - PositiveMap.fold add_assoclist nbassoc assoc. +Definition merge_assocset (nbassoc assoc : assocset) : assocset := + PositiveMap.fold add_assocset nbassoc assoc. -Definition empty_assoclist : assoclist := PositiveMap.empty value. +Definition empty_assocset : assocset := PositiveMap.empty value. (*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_assoclist nbassoc assoc) empty_assoclist f c) + OK (State m (merge_assocset nbassoc assoc) empty_assocset f c) | Error msg => Error msg | _ => Error (msg "Verilog: Wrong state") end.*) -(*Fixpoint mi_run (f : fextclk) (assoc : assoclist) (m : list module_item) (n : nat) - {struct n} : res assoclist := +(*Fixpoint mi_run (f : fextclk) (assoc : assocset) (m : list module_item) (n : nat) + {struct n} : res assocset := match n with | S n' => do assoc' <- mi_run f assoc m n'; - match mi_step_commit (State assoc' empty_assoclist f (Pos.of_nat n')) m with + match mi_step_commit (State assoc' empty_assocset f (Pos.of_nat n')) m with | OK (State assoc _ _ _) => OK assoc | Error m => Error m end @@ -553,12 +553,12 @@ other arguments to the module are also to be supported. *) Definition initial_fextclk (m : module) : fextclk := fun x => match x with - | S O => (add_assoclist (fst (mod_reset m)) (ZToValue 1 1) empty_assoclist) - | _ => (add_assoclist (fst (mod_reset m)) (ZToValue 1 0) empty_assoclist) + | S O => (add_assocset (fst (mod_reset m)) (ZToValue 1 1) empty_assocset) + | _ => (add_assocset (fst (mod_reset m)) (ZToValue 1 0) empty_assocset) end. -(*Definition module_run (n : nat) (m : module) : res assoclist := - mi_run (initial_fextclk m) empty_assoclist (mod_body m) n.*) +(*Definition module_run (n : nat) (m : module) : res assocset := + mi_run (initial_fextclk m) empty_assocset (mod_body m) n.*) Local Close Scope error_monad_scope. @@ -606,18 +606,18 @@ Definition genv := Genv.t unit unit. Inductive step : state -> state -> Prop := | step_module : forall m stvar stvar' cycle f assoc0 assoc1 assoc2 nbassoc, - mis_stepp (State m assoc0 empty_assoclist f cycle stvar) + mis_stepp (State m assoc0 empty_assocset f cycle stvar) m.(mod_body) (State m assoc1 nbassoc f cycle stvar) -> - assoc2 = merge_assoclist nbassoc assoc1 -> + assoc2 = merge_assocset nbassoc assoc1 -> Some stvar' = assoc2!(fst m.(mod_state)) -> - step (State m assoc0 empty_assoclist f cycle stvar) - (State m assoc2 empty_assoclist f (S cycle) stvar') + step (State m assoc0 empty_assocset f cycle stvar) + (State m assoc2 empty_assocset f (S cycle) stvar') | step_finish : forall m assoc f cycle stvar result, assoc!(fst m.(mod_finish)) = Some (1'h"1") -> assoc!(fst m.(mod_return)) = Some result -> - step (State m assoc empty_assoclist f cycle stvar) + step (State m assoc empty_assocset f cycle stvar) (Finishstate result). Hint Constructors step : verilog. @@ -625,7 +625,7 @@ Hint Constructors step : verilog. | initial_state_intro: forall hmi tmi, hmi::tmi = mod_body m -> - initial_state m (State m hmi tmi empty_assoclist empty_assoclist (initial_fextclk m) O xH). + initial_state m (State m hmi tmi empty_assocset empty_assocset (initial_fextclk m) O xH). (** A final state is a [Returnstate] with an empty call stack. *) -- cgit From fa39e09d403cfba1b1e6c359362e54600dfc28b0 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 7 May 2020 23:12:59 +0100 Subject: Use associations instead of state --- src/verilog/Value.v | 4 ++ src/verilog/Verilog.v | 135 ++++++++++++++++++++++++-------------------------- 2 files changed, 69 insertions(+), 70 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 8a0716d..a96d67c 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -200,6 +200,10 @@ Proof. split; intros. unfold valueeqb in H. destruct (value_eq_size v1 v2) eqn:?. - destruct v1, v2. simpl in H. +Abort. + +Definition value_int_eqb (v : value) (i : int) : bool := + Z.eqb (valueToZ v) (Int.unsigned i). (** Arithmetic operations over [value], interpreting them as signed or unsigned depending on the operation. diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index a35fb58..ba7e83d 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -41,6 +41,12 @@ Definition szreg : Type := reg * nat. Definition assocset : Type := PositiveMap.t value. +Record associations : Type := + mkassociations { + assoc_blocking : assocset; + assoc_nonblocking : assocset + }. + (** * Verilog AST The Verilog AST is defined here, which is the target language of the @@ -409,72 +415,61 @@ Fixpoint stmnt_run' (n : nat) (s : state) (st : stmnt) {struct n} : res state := Definition stmnt_run (s : state) (st : stmnt) : res state := stmnt_run' (stmnt_height st) s st. *) -Inductive state_fext_assoc : state -> fext * assocset -> Prop := -| get_state_fext_assoc : - forall c assoc nbassoc f cycle pc, - state_fext_assoc (State c assoc nbassoc f cycle pc) (f cycle, assoc). -Hint Constructors state_fext_assoc : verilog. - -Inductive stmnt_runp: state -> stmnt -> state -> Prop := +Inductive stmnt_runp: fext -> associations -> stmnt -> associations -> Prop := | stmnt_runp_Vskip: - forall s, stmnt_runp s Vskip s + forall f a, stmnt_runp f a Vskip a | stmnt_runp_Vseq: - forall st1 st2 s0 s1 s2, - stmnt_runp s0 st1 s1 -> - stmnt_runp s1 st2 s2 -> - stmnt_runp s0 (Vseq st1 st2) s2 + forall f st1 st2 as0 as1 as2, + stmnt_runp f as0 st1 as1 -> + stmnt_runp f as1 st2 as2 -> + stmnt_runp f as0 (Vseq st1 st2) as2 | stmnt_runp_Vcond_true: - forall s0 f assoc s1 c vc stt stf, - state_fext_assoc s0 (f, assoc) -> - expr_runp f assoc c vc -> + forall as0 as1 f c vc stt stf, + expr_runp f as0.(assoc_blocking) c vc -> valueToBool vc = true -> - stmnt_runp s0 stt s1 -> - stmnt_runp s0 (Vcond c stt stf) s1 + stmnt_runp f as0 stt as1 -> + stmnt_runp f as0 (Vcond c stt stf) as1 | stmnt_runp_Vcond_false: - forall s0 f assoc s1 c vc stt stf, - state_fext_assoc s0 (f, assoc) -> - expr_runp f assoc c vc -> + forall as0 as1 f c vc stt stf, + expr_runp f as0.(assoc_blocking) c vc -> valueToBool vc = false -> - stmnt_runp s0 stf s1 -> - stmnt_runp s0 (Vcond c stt stf) s1 + stmnt_runp f as0 stf as1 -> + stmnt_runp f as0 (Vcond c stt stf) as1 | stmnt_runp_Vcase_nomatch: - forall e ve s0 f assoc s1 me mve sc cs def, - state_fext_assoc s0 (f, assoc) -> - expr_runp f assoc e ve -> - expr_runp f assoc me mve -> + forall e ve as0 f as1 me mve sc cs def, + expr_runp f as0.(assoc_blocking) e ve -> + expr_runp f as0.(assoc_blocking) me mve -> mve <> ve -> - stmnt_runp s0 (Vcase e cs def) s1 -> - stmnt_runp s0 (Vcase e ((me, sc)::cs) def) s1 + stmnt_runp f as0 (Vcase e cs def) as1 -> + stmnt_runp f as0 (Vcase e ((me, sc)::cs) def) as1 | stmnt_runp_Vcase_match: - forall e ve s0 f assoc s1 me mve sc cs def, - state_fext_assoc s0 (f, assoc) -> - expr_runp f assoc e ve -> - expr_runp f assoc me mve -> + forall e ve as0 f as1 me mve sc cs def, + expr_runp f as0.(assoc_blocking) e ve -> + expr_runp f as0.(assoc_blocking) me mve -> mve = ve -> - stmnt_runp s0 sc s1 -> - stmnt_runp s0 (Vcase e ((me, sc)::cs) def) s1 + stmnt_runp f as0 sc as1 -> + stmnt_runp f as0 (Vcase e ((me, sc)::cs) def) as1 | stmnt_runp_Vcase_default: - forall s0 f assoc st s1 e ve, - state_fext_assoc s0 (f, assoc) -> - expr_runp f assoc e ve -> - stmnt_runp s0 st s1 -> - stmnt_runp s0 (Vcase e nil (Some st)) s1 + forall as0 as1 f st e ve, + expr_runp f as0.(assoc_blocking) e ve -> + stmnt_runp f as0 st as1 -> + stmnt_runp f as0 (Vcase e nil (Some st)) as1 | stmnt_runp_Vblock: - forall lhs name rhs rhsval c assoc assoc' nbassoc f cycle pc, + forall lhs name rhs rhsval assoc assoc' nbassoc f, assign_name lhs = OK name -> - expr_runp (f cycle) assoc rhs rhsval -> + expr_runp f assoc rhs rhsval -> assoc' = (PositiveMap.add name rhsval assoc) -> - stmnt_runp (State c assoc nbassoc f cycle pc) - (Vblock lhs rhs) - (State c assoc' nbassoc f cycle pc) + stmnt_runp f (mkassociations assoc nbassoc) + (Vblock lhs rhs) + (mkassociations assoc nbassoc) | stmnt_runp_Vnonblock: - forall lhs name rhs rhsval c assoc nbassoc nbassoc' f cycle pc, + forall lhs name rhs rhsval assoc nbassoc nbassoc' f, assign_name lhs = OK name -> - expr_runp (f cycle) assoc rhs rhsval -> + expr_runp f assoc rhs rhsval -> nbassoc' = (PositiveMap.add name rhsval nbassoc) -> - stmnt_runp (State c assoc nbassoc f cycle pc) - (Vnonblock lhs rhs) - (State c assoc nbassoc' f cycle pc). + stmnt_runp f (mkassociations assoc nbassoc) + (Vnonblock lhs rhs) + (mkassociations assoc nbassoc'). Hint Constructors stmnt_runp : verilog. (*Fixpoint mi_step (s : state) (m : list module_item) {struct m} : res state := @@ -490,33 +485,33 @@ Hint Constructors stmnt_runp : verilog. | nil => OK s end.*) -Inductive mi_stepp : state -> module_item -> state -> Prop := +Inductive mi_stepp : fext -> associations -> module_item -> associations -> Prop := | mi_stepp_Valways : - forall s0 st s1 c, - stmnt_runp s0 st s1 -> - mi_stepp s0 (Valways c st) s1 + forall f s0 st s1 c, + stmnt_runp f s0 st s1 -> + mi_stepp f s0 (Valways c st) s1 | mi_stepp_Valways_ff : - forall s0 st s1 c, - stmnt_runp s0 st s1 -> - mi_stepp s0 (Valways_ff c st) s1 + forall f s0 st s1 c, + stmnt_runp f s0 st s1 -> + mi_stepp f s0 (Valways_ff c st) s1 | mi_stepp_Valways_comb : - forall s0 st s1 c, - stmnt_runp s0 st s1 -> - mi_stepp s0 (Valways_comb c st) s1 + forall f s0 st s1 c, + stmnt_runp f s0 st s1 -> + mi_stepp f s0 (Valways_comb c st) s1 | mi_stepp_Vdecl : - forall s lhs rhs, - mi_stepp s (Vdecl lhs rhs) s. + forall f s lhs rhs, + mi_stepp f s (Vdecl lhs rhs) s. Hint Constructors mi_stepp : verilog. -Inductive mis_stepp : state -> list module_item -> state -> Prop := +Inductive mis_stepp : fext -> associations -> list module_item -> associations -> Prop := | mis_stepp_Cons : - forall mi mis s0 s1 s2, - mi_stepp s0 mi s1 -> - mis_stepp s1 mis s2 -> - mis_stepp s0 (mi :: mis) s2 + forall f mi mis s0 s1 s2, + mi_stepp f s0 mi s1 -> + mis_stepp f s1 mis s2 -> + mis_stepp f s0 (mi :: mis) s2 | mis_stepp_Nil : - forall s, - mis_stepp s nil s. + forall f s, + mis_stepp f s nil s. Hint Constructors mis_stepp : verilog. Definition add_assocset (r : reg) (v : value) (assoc : assocset) : assocset := @@ -606,9 +601,9 @@ Definition genv := Genv.t unit unit. Inductive step : state -> state -> Prop := | step_module : forall m stvar stvar' cycle f assoc0 assoc1 assoc2 nbassoc, - mis_stepp (State m assoc0 empty_assocset f cycle stvar) + mis_stepp (f cycle) (mkassociations assoc0 empty_assocset) m.(mod_body) - (State m assoc1 nbassoc f cycle stvar) -> + (mkassociations assoc1 nbassoc) -> assoc2 = merge_assocset nbassoc assoc1 -> Some stvar' = assoc2!(fst m.(mod_state)) -> step (State m assoc0 empty_assocset f cycle stvar) -- cgit From 4bdabd1828c48e74c6b1f701f57a1b3c421a95fb Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 7 May 2020 23:13:25 +0100 Subject: Redefine HTL for intermediate Verilog language --- src/translation/Veriloggenspec.v | 18 +++++ src/verilog/HTL.v | 145 +++++++++++++++++++-------------------- 2 files changed, 87 insertions(+), 76 deletions(-) diff --git a/src/translation/Veriloggenspec.v b/src/translation/Veriloggenspec.v index 67e2cad..09d08ed 100644 --- a/src/translation/Veriloggenspec.v +++ b/src/translation/Veriloggenspec.v @@ -1,3 +1,21 @@ +(* + * CoqUp: 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 . + *) + From Coq Require Import FSets.FMapPositive. From compcert Require RTL Op Maps. From coqup Require Import Coquplib Verilog Veriloggen Value. diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v index 1d156ad..3d863b2 100644 --- a/src/verilog/HTL.v +++ b/src/verilog/HTL.v @@ -16,89 +16,82 @@ * along with this program. If not, see . *) -(** 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. - *) - -From coqup.common Require Import Coquplib. - -From compcert Require Import Maps. -From compcert Require Op AST Memory Registers. - -Definition node := positive. - -Definition reg := Registers.reg. +From Coq Require Import FSets.FMapPositive. +From coqup Require Import Coquplib Value. +From coqup Require Verilog. +From compcert Require Events Globalenvs Smallstep Integers. -Definition ident := AST.ident. +Import HexNotationValue. -Inductive instruction : Type := -| Hnop : node -> instruction -| Hnonblock : Op.operation -> list reg -> reg -> node -> instruction - (** [Hnonblock op args res next] Defines a nonblocking assignment to a - register, using the operation defined in Op.v. *) -| Hload : AST.memory_chunk -> Op.addressing -> list reg -> reg -> node -> instruction -| Hstore : AST.memory_chunk -> Op.addressing -> list reg -> reg -> node -> instruction -| Hinst : AST.signature -> ident -> reg -> node -> instruction - (** [Hinst sig fun args rst strt end res next] Defines the start of a - module instantiation, meaning the function will run until the result is - returned. *) -| Htailcall : AST.signature -> ident -> list reg -> instruction -| Hcond : Op.condition -> list reg -> node -> node -> instruction -| Hjumptable : reg -> list node -> instruction -| Hfinish : option reg -> instruction. - -Record inst : Type := - mkinst { - inst_moddecl : ident; - inst_args : list reg - }. +(** 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. *) -Definition code : Type := PTree.t instruction. +Notation "a ! b" := (PositiveMap.find b a) (at level 1). -Definition instances : Type := PTree.t inst. +Definition reg := positive. +Definition node := positive. -Definition empty_instances : instances := PTree.empty inst. +Definition datapath := PositiveMap.t Verilog.stmnt. +Definition controllogic := PositiveMap.t Verilog.stmnt. -(** Function declaration for VTL also contain a construction which describes the - functions that are called in the current function. This information is used - to print out *) -Record module : Type := +Record module: Type := mkmodule { - mod_sig : AST.signature; mod_params : list reg; - mod_stacksize : Z; - mod_code : code; - mod_insts : instances; - mod_entrypoint : node + mod_datapath : datapath; + mod_controllogic : controllogic; + mod_entrypoint : node; + mod_st : reg; + mod_finish : reg; + mod_return : reg }. -Definition moddecl := AST.fundef module. - -Definition design := AST.program moddecl unit. - -Definition modsig (md : moddecl) := - match md with - | AST.Internal m => mod_sig m - | AST.External ef => AST.ef_sig ef - end. - -(** Describes various transformations that can be applied to HTL. This applies - the transformation to each instruction in the function and returns the new - function with the modified instructions. *) -Section TRANSF. - - Variable transf_instr : node -> instruction -> instruction. - - Definition transf_module (m : module) : module := - mkmodule - m.(mod_sig) - m.(mod_params) - m.(mod_stacksize) - (PTree.map transf_instr m.(mod_code)) - m.(mod_insts) - m.(mod_entrypoint). - -End TRANSF. +(** * Operational Semantics *) + +Definition genv := Globalenvs.Genv.t unit unit. +Definition genv_empty := Globalenvs.Genv.empty_genv unit unit nil. + +Inductive state : Type := +| State : + forall (m : module) + (st : node) + (assoc : Verilog.assocset), + state +| Returnstate : forall v : value, state. + +Inductive step : genv -> state -> Events.trace -> state -> Prop := +| step_module : + forall g t m st ctrl data assoc0 assoc1 assoc2 assoc3 nbassoc0 nbassoc1 f, + m.(mod_controllogic)!st = Some ctrl -> + m.(mod_datapath)!st = Some data -> + Verilog.stmnt_runp f + (Verilog.mkassociations assoc0 Verilog.empty_assocset) + ctrl + (Verilog.mkassociations assoc1 nbassoc0) -> + Verilog.stmnt_runp f + (Verilog.mkassociations assoc1 nbassoc0) + data + (Verilog.mkassociations assoc2 nbassoc1) -> + assoc3 = Verilog.merge_assocset nbassoc1 assoc2 -> + step g (State m st assoc0) t (State m st assoc3) +| step_finish : + forall g t m st assoc retval, + assoc!(m.(mod_finish)) = Some (1'h"1") -> + assoc!(m.(mod_return)) = Some retval -> + step g (State m st assoc) t (Returnstate retval). +Hint Constructors step : htl. + +Inductive initial_state (m : module) : state -> Prop := +| initial_state_intro : forall st, + st = m.(mod_entrypoint) -> + initial_state m (State m st Verilog.empty_assocset). + +Inductive final_state : state -> Integers.int -> Prop := +| final_state_intro : forall retval retvali, + value_int_eqb retval retvali = true -> + final_state (Returnstate retval) retvali. + +Definition semantics (m : module) := + Smallstep.Semantics step (initial_state m) final_state genv_empty. -- cgit From 78a7b735373c0607f66d863b224cb6bb7742d1f1 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 7 May 2020 23:13:54 +0100 Subject: Remove HTLgen and create the specification --- src/translation/HTLgen.v | 163 ------------------------------------------- src/translation/HTLgenspec.v | 92 ++++++++++++++++++++++++ 2 files changed, 92 insertions(+), 163 deletions(-) create mode 100644 src/translation/HTLgenspec.v diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index d7f88c1..8eac48c 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -17,170 +17,7 @@ *) From coqup Require Import HTL Coquplib. - From compcert Require Import Maps. From compcert Require Errors. From compcert Require Import AST RTL. -Record state: Type := mkstate { - st_nextinst: positive; - st_instances: instances; -}. - -Inductive res (A: Type) (S: Type): Type := - | Error: Errors.errmsg -> res A S - | OK: A -> S -> res A S. - -Arguments OK [A S]. -Arguments Error [A S]. - -Definition mon (A: Type) : Type := res A state. - -Definition ret {A: Type} (x: A) (s: state) : mon A := OK x s. - -Definition bind {A B: Type} (f: mon A) (g: A -> mon B) : mon B := - match f with - | Error msg => Error msg - | OK a s => g a - end. - -Definition bind2 {A B C: Type} (f: mon (A * B)) (g: A -> B -> mon C) : mon C := - bind f (fun xy => g (fst xy) (snd xy)). - -Definition bindS {A B: Type} (f: mon A) (g: A -> state -> mon B) : mon B := - match f with - | Error msg => Error msg - | OK a s => g a s - end. - -Notation "'do' X <- A ; B" := (bind A (fun X => B)) - (at level 200, X ident, A at level 100, B at level 200). -Notation "'do' ( X , Y ) <- A ; B" := (bindS A (fun X Y => B)) - (at level 200, X ident, Y ident, A at level 100, B at level 200). - -Definition handle_error {A: Type} (f g: mon A) : mon A := - match f with - | OK a s' => OK a s' - | Error _ => g - end. - -Definition init_state : state := - mkstate 1%positive empty_instances. - -Module PTree. - Export Maps.PTree. - - Fixpoint xtraverse {A B : Type} (f : positive -> A -> state -> mon B) - (m : PTree.t A) (s : state) (i : positive) - {struct m} : mon (PTree.t B) := - match m with - | Leaf => OK Leaf s - | Node l o r => - let newo := - match o with - | None => OK None s - | Some x => - match f (prev i) x s with - | Error err => Error err - | OK val s' => OK (Some val) s' - end - end in - match newo with - | OK no s => - do (nl, s') <- xtraverse f l s (xO i); - do (nr, s'') <- xtraverse f r s' (xI i); - OK (Node nl no nr) s'' - | Error msg => Error msg - end - end. - - Definition traverse {A B : Type} (f : positive -> A -> state -> mon B) m := - xtraverse f m init_state xH. - - Definition traverse1 {A B : Type} (f : A -> state -> mon B) := traverse (fun _ => f). - -End PTree. - -Definition transf_instr (pc: node) (rtl: RTL.instruction) (s: state) - : mon HTL.instruction := - match rtl with - | RTL.Inop n => - (** Nop instruction should just become a Skip in Verilog, which is just a - semicolon. *) - ret (HTL.Hnop n) s - | RTL.Iop op args res n => - (** Perform an arithmetic operation over registers. This will just become - one binary operation in Verilog. This will contain a list of registers, - so these will just become a chain of binary operations in Verilog. *) - ret (HTL.Hnonblock op args res n) s - | RTL.Iload chunk addr args dst n => - (** Load from memory, which will maybe become a load from RAM, however, I'm - not sure yet how memory will be implemented or how it will formalised - be. *) - ret (HTL.Hload chunk addr args dst n) s - | RTL.Istore chunk addr args src n => - (** How memory will be laid out will also affect how stores are handled. For - now hopefully these can be ignored, and hopefull they are not used that - often when they are not required in C. *) - ret (HTL.Hstore chunk addr args src n) s - | RTL.Icall sig ros args res n => - (** Function call should become a module instantiation with start and end - signals appropriately wired up. *) - match ros with - | inr i => - let inst := mkinst i args in - let newinstances := PTree.set s.(st_nextinst) inst s.(st_instances) in - ret (HTL.Hinst sig i res n) - (mkstate ((s.(st_nextinst) + 1)%positive) - newinstances) - | _ => Error (Errors.msg "Function pointers are not supported.") - end - | RTL.Itailcall sig ros args => - (** Should be converted into a reset of the modules state, setting the - initial variables correctly. This would simulate a tailcall as it is - similar to jumping back to the top of the function in assembly. *) - match ros with - | inr i => ret (HTL.Htailcall sig i args) s - | _ => Error (Errors.msg "Function pointers are not supported.") - end - | RTL.Ibuiltin ef args res n => - (** Builtin functions will have to supported manually, by implementing the - Verilog myself. *) - Error (Errors.msg "builtin functions are not supported.") - | RTL.Icond cond args n1 n2 => - (** Will be converted into two separate processes that are combined by a mux - at the end, with a start flag that propagates in the construct that should - be chosen. *) - ret (HTL.Hcond cond args n1 n2) s - | RTL.Ijumptable arg tbl => - (** A jump to a specific instruction in the list, this will require setting - the state in the state machine appropriately. This is trivial to - implement if no scheduling is involved, but as soon as that isn't the - case it might be difficult to find that location. It would help to - transform the CFG to a few basic blocks first which cannot have any - branching. *) - ret (HTL.Hjumptable arg tbl) s - | RTL.Ireturn r => - (** The return value of the function, which will just mean that the finished - signal goes high and the result is stored in the result register. *) - ret (HTL.Hfinish r) s - end. - -Definition transf_function (f: function) : Errors.res module := - match (PTree.traverse transf_instr f.(fn_code)) with - | OK code s => - Errors.OK (mkmodule - f.(fn_sig) - f.(fn_params) - f.(fn_stacksize) - code - s.(st_instances) - f.(fn_entrypoint)) - | Error err => Errors.Error err - end. - -Definition transf_fundef (fd: fundef) : Errors.res moddecl := - transf_partial_fundef transf_function fd. - -Definition transf_program (p: program) : Errors.res design := - transform_partial_program transf_fundef p. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v new file mode 100644 index 0000000..fb1c70d --- /dev/null +++ b/src/translation/HTLgenspec.v @@ -0,0 +1,92 @@ +(* + * CoqUp: 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 . + *) + +From Coq Require Import FSets.FMapPositive. +From compcert Require RTL Op Maps. +From coqup Require Import Coquplib Verilog Veriloggen Value HTL. + +(** * Relational specification of the translation *) + +(** 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 *) + +Inductive tr_op : Op.operation -> list reg -> expr -> Prop := +| tr_op_Omove : forall r, tr_op Op.Omove (r::nil) (Vvar r) +| tr_op_Ointconst : forall n l, tr_op (Op.Ointconst n) l (Vlit (intToValue n)) +| tr_op_Oneg : forall r, tr_op Op.Oneg (r::nil) (Vunop Vneg (Vvar r)) +| tr_op_Osub : forall r1 r2, tr_op Op.Osub (r1::r2::nil) (bop Vsub r1 r2) +| tr_op_Omul : forall r1 r2, tr_op Op.Omul (r1::r2::nil) (bop Vmul r1 r2) +| tr_op_Omulimm : forall r n, tr_op (Op.Omulimm n) (r::nil) (boplit Vmul r n) +| tr_op_Odiv : forall r1 r2, tr_op Op.Odiv (r1::r2::nil) (bop Vdiv r1 r2) +| tr_op_Odivu : forall r1 r2, tr_op Op.Odivu (r1::r2::nil) (bop Vdivu r1 r2) +| tr_op_Omod : forall r1 r2, tr_op Op.Omod (r1::r2::nil) (bop Vmod r1 r2) +| tr_op_Omodu : forall r1 r2, tr_op Op.Omodu (r1::r2::nil) (bop Vmodu r1 r2) +| tr_op_Oand : forall r1 r2, tr_op Op.Oand (r1::r2::nil) (bop Vand r1 r2) +| tr_op_Oandimm : forall n r, tr_op (Op.Oandimm n) (r::nil) (boplit Vand r n) +| tr_op_Oor : forall r1 r2, tr_op Op.Oor (r1::r2::nil) (bop Vor r1 r2) +| tr_op_Oorimm : forall n r, tr_op (Op.Oorimm n) (r::nil) (boplit Vor r n) +| tr_op_Oxor : forall r1 r2, tr_op Op.Oxor (r1::r2::nil) (bop Vxor r1 r2) +| tr_op_Oxorimm : forall n r, tr_op (Op.Oxorimm n) (r::nil) (boplit Vxor r n) +| tr_op_Onot : forall r, tr_op Op.Onot (r::nil) (Vunop Vnot (Vvar r)) +| tr_op_Oshl : forall r1 r2, tr_op Op.Oshl (r1::r2::nil) (bop Vshl r1 r2) +| tr_op_Oshlimm : forall n r, tr_op (Op.Oshlimm n) (r::nil) (boplit Vshl r n) +| tr_op_Oshr : forall r1 r2, tr_op Op.Oshr (r1::r2::nil) (bop Vshr r1 r2) +| tr_op_Oshrimm : forall n r, tr_op (Op.Oshrimm n) (r::nil) (boplit Vshr r n) +| tr_op_Ocmp : forall c l e s s' i, translate_condition c l s = OK e s' i -> tr_op (Op.Ocmp c) l e +| tr_op_Olea : forall a l e s s' i, translate_eff_addressing a l s = OK e s' i -> tr_op (Op.Olea a) l e. + +Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Prop := +| tr_instr_Inop : + forall n, + tr_instr fin rtrn st (RTL.Inop n) Vskip (state_goto st n) +| tr_instr_Iop : + forall n op args e dst, + tr_op op args e -> + tr_instr fin rtrn st (RTL.Iop op args dst n) (Vnonblock (Vvar dst) e) (state_goto st n) +| tr_instr_Icond : + forall n1 n2 cond args s s' i c, + translate_condition cond args s = OK c s' i -> + tr_instr fin rtrn st (RTL.Icond cond args n1 n2) Vskip (state_cond st c n1 n2) +| tr_instr_Ireturn_None : + tr_instr fin rtrn st (RTL.Ireturn None) (block fin (Vlit (ZToValue 1%nat 1%Z))) Vskip +| tr_instr_Ireturn_Some : + forall r, + tr_instr fin rtrn st (RTL.Ireturn (Some r)) + (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r))) Vskip. + +Inductive tr_code (c : RTL.code) (stmnts trans : PositiveMap.t stmnt) + (fin rtrn st : reg) : Prop := + tr_code_intro : + forall i s t n, + Maps.PTree.get n c = Some i -> + stmnts!n = Some s -> + trans!n = Some t -> + tr_instr fin rtrn st i s t -> + tr_code c stmnts trans fin rtrn st. + +Inductive tr_module (f : RTL.function) : module -> Prop := + tr_module_intro : + forall c data control fin rtrn st, + tr_code c data control fin rtrn st -> + tr_module f (mkmodule + f.(RTL.fn_params) + data + control + f.(RTL.fn_entrypoint) + st fin rtrn). -- cgit From 32f0f542c18b73303613b53573e6c61bc7ae6890 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 7 May 2020 23:14:48 +0100 Subject: Add match_states Inductive --- src/translation/HTLgenproof.v | 29 +++++++++++++++++++++++++++++ 1 file changed, 29 insertions(+) create mode 100644 src/translation/HTLgenproof.v diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v new file mode 100644 index 0000000..898fafd --- /dev/null +++ b/src/translation/HTLgenproof.v @@ -0,0 +1,29 @@ +(* + * CoqUp: 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 . + *) + +From coqup Require Import HTLgenspec. +From coqup Require HTL Verilog. +From compcert Require RTL. + +Parameter match_assocset : RTL.regset -> Verilog.assocset -> Prop. + +Inductive match_states : RTL.state -> HTL.state -> Prop := +| match_state : forall (rs : RTL.regset) assoc sf f sp rs mem m st, + match_assocset rs assoc -> + match_states (RTL.State sf f sp st rs mem) + (HTL.State m st assoc). -- cgit From 32990af23915c15e5cfd4cfe32c47cafa508f11d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 8 May 2020 23:34:15 +0100 Subject: Add AssocMap --- src/verilog/AssocMap.v | 60 ++++++++++++++++++++++++++++++++++++++++++++++ src/verilog/HTL.v | 10 ++++---- src/verilog/Test.v | 6 ++--- src/verilog/Verilog.v | 64 ++++++++++++++++++++------------------------------ 4 files changed, 93 insertions(+), 47 deletions(-) create mode 100644 src/verilog/AssocMap.v diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v new file mode 100644 index 0000000..e550072 --- /dev/null +++ b/src/verilog/AssocMap.v @@ -0,0 +1,60 @@ +(* + * CoqUp: 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 . + *) + +From coqup Require Import Coquplib Value. +From Coq Require Import FSets.FMapPositive. + +Definition reg := positive. + +Module AssocMap := PositiveMap. + +Module AssocMapExt. + Import AssocMap. + + Section Operations. + + Variable elt : Type. + + Definition merge : t elt -> t elt -> t elt := fold (@add elt). + + Definition find_default (a : elt) (k : reg) (m : t elt) : elt := + match find k m with + | None => a + | Some v => v + end. + + End Operations. + +End AssocMapExt. +Import AssocMapExt. + +Definition assocmap := AssocMap.t value. + +Definition find_assocmap (n : nat) : reg -> assocmap -> value := + find_default value (ZToValue n 0). + +Definition empty_assocmap : assocmap := AssocMap.empty value. + +Definition merge_assocmap : assocmap -> assocmap -> assocmap := merge value. + +Module AssocMapNotation. + Notation "a ! b" := (AssocMap.find b a) (at level 1). + Notation "a # ( b , c )" := (find_assocmap c b a) (at level 1). + Notation "a # b" := (find_assocmap 32 b a) (at level 1). + Notation "a ## b" := (List.map (fun c => find_assocmap 32 c a) b) (at level 1). +End AssocMapNotation. diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v index 3d863b2..147d24b 100644 --- a/src/verilog/HTL.v +++ b/src/verilog/HTL.v @@ -17,7 +17,7 @@ *) From Coq Require Import FSets.FMapPositive. -From coqup Require Import Coquplib Value. +From coqup Require Import Coquplib Value AssocMap. From coqup Require Verilog. From compcert Require Events Globalenvs Smallstep Integers. @@ -57,7 +57,7 @@ Inductive state : Type := | State : forall (m : module) (st : node) - (assoc : Verilog.assocset), + (assoc : assocmap), state | Returnstate : forall v : value, state. @@ -67,14 +67,14 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := m.(mod_controllogic)!st = Some ctrl -> m.(mod_datapath)!st = Some data -> Verilog.stmnt_runp f - (Verilog.mkassociations assoc0 Verilog.empty_assocset) + (Verilog.mkassociations assoc0 empty_assocmap) ctrl (Verilog.mkassociations assoc1 nbassoc0) -> Verilog.stmnt_runp f (Verilog.mkassociations assoc1 nbassoc0) data (Verilog.mkassociations assoc2 nbassoc1) -> - assoc3 = Verilog.merge_assocset nbassoc1 assoc2 -> + assoc3 = merge_assocmap nbassoc1 assoc2 -> step g (State m st assoc0) t (State m st assoc3) | step_finish : forall g t m st assoc retval, @@ -86,7 +86,7 @@ Hint Constructors step : htl. Inductive initial_state (m : module) : state -> Prop := | initial_state_intro : forall st, st = m.(mod_entrypoint) -> - initial_state m (State m st Verilog.empty_assocset). + initial_state m (State m st empty_assocmap). Inductive final_state : state -> Integers.int -> Prop := | final_state_intro : forall retval retvali, diff --git a/src/verilog/Test.v b/src/verilog/Test.v index f24612e..90c5312 100644 --- a/src/verilog/Test.v +++ b/src/verilog/Test.v @@ -88,10 +88,10 @@ Proof. trivial. Qed. Definition test_fextclk := initial_fextclk test_output_module. Lemma manual_simulation : - step (State test_output_module empty_assocset empty_assocset + step (State test_output_module empty_assocmap empty_assocmap test_fextclk 1 (32'h"1")) - (State test_output_module (add_assocset 7%positive (32'h"3") empty_assocset) - empty_assocset test_fextclk 2 (32'h"3")). + (State test_output_module (add_assocmap 7%positive (32'h"3") empty_assocmap) + empty_assocmap test_fextclk 2 (32'h"3")). Proof. repeat (econstructor; eauto); try (unfold ZToValue; instantiate (1 := eq_refl (vsize (1'h"1"))); auto); diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index ba7e83d..c0bce25 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -27,24 +27,21 @@ From Coq Require Import Import ListNotations. -From coqup Require Import common.Coquplib common.Show verilog.Value. +From coqup Require Import common.Coquplib common.Show verilog.Value AssocMap. From compcert Require Integers Events. From compcert Require Import Errors Smallstep Globalenvs. Import HexNotationValue. - -Notation "a ! b" := (PositiveMap.find b a) (at level 1). +Import AssocMapNotation. Definition reg : Type := positive. Definition node : Type := positive. Definition szreg : Type := reg * nat. -Definition assocset : Type := PositiveMap.t value. - Record associations : Type := mkassociations { - assoc_blocking : assocset; - assoc_nonblocking : assocset + assoc_blocking : assocmap; + assoc_nonblocking : assocmap }. (** * Verilog AST @@ -58,8 +55,6 @@ 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. *) -Definition estate : Type := assocset * assocset. - (** ** Binary Operators These are the binary operations that can be represented in Verilog. Ideally, @@ -167,8 +162,7 @@ Record module : Type := mkmodule { mod_state : szreg; (**r Variable that defines the current state, it should be internal. *) mod_args : list szreg; mod_body : list module_item; -}. - + }. (** Convert a [positive] to an expression directly, setting it to the right size. *) Definition posToLit (p : positive) : expr := @@ -204,8 +198,8 @@ retrieved and set appropriately. Inductive state : Type := | State: forall (m : module) - (assoc : assocset) - (nbassoc : assocset) + (assoc : assocmap) + (nbassoc : assocmap) (f : fextclk) (cycle : nat) (stvar : value), @@ -246,7 +240,7 @@ Definition unop_run (op : unop) : value -> value := | Vnot => vbitneg end. -Inductive expr_runp : fext -> assocset -> expr -> value -> Prop := +Inductive expr_runp : fext -> assocmap -> expr -> value -> Prop := | erun_Vlit : forall fext assoc v, expr_runp fext assoc (Vlit v) v @@ -312,7 +306,7 @@ Definition access_fext (f : fext) (r : reg) : res value := end. (* TODO FIX Vvar case without default *) -(*Fixpoint expr_run (assoc : assocset) (e : expr) +(*Fixpoint expr_run (assoc : assocmap) (e : expr) {struct e} : res value := match e with | Vlit v => OK v @@ -461,7 +455,7 @@ Inductive stmnt_runp: fext -> associations -> stmnt -> associations -> Prop := assoc' = (PositiveMap.add name rhsval assoc) -> stmnt_runp f (mkassociations assoc nbassoc) (Vblock lhs rhs) - (mkassociations assoc nbassoc) + (mkassociations assoc' nbassoc) | stmnt_runp_Vnonblock: forall lhs name rhs rhsval assoc nbassoc nbassoc' f, assign_name lhs = OK name -> @@ -514,28 +508,20 @@ Inductive mis_stepp : fext -> associations -> list module_item -> associations - mis_stepp f s nil s. Hint Constructors mis_stepp : verilog. -Definition add_assocset (r : reg) (v : value) (assoc : assocset) : assocset := - PositiveMap.add r v assoc. - -Definition merge_assocset (nbassoc assoc : assocset) : assocset := - PositiveMap.fold add_assocset nbassoc assoc. - -Definition empty_assocset : assocset := PositiveMap.empty value. - (*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_assocset nbassoc assoc) empty_assocset f c) + OK (State m (merge_assocmap nbassoc assoc) empty_assocmap f c) | Error msg => Error msg | _ => Error (msg "Verilog: Wrong state") - end.*) + end. -(*Fixpoint mi_run (f : fextclk) (assoc : assocset) (m : list module_item) (n : nat) - {struct n} : res assocset := +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_assocset f (Pos.of_nat n')) m with + 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 @@ -548,12 +534,12 @@ other arguments to the module are also to be supported. *) Definition initial_fextclk (m : module) : fextclk := fun x => match x with - | S O => (add_assocset (fst (mod_reset m)) (ZToValue 1 1) empty_assocset) - | _ => (add_assocset (fst (mod_reset m)) (ZToValue 1 0) empty_assocset) + | S O => (AssocMap.add (fst (mod_reset m)) (ZToValue 1 1) empty_assocmap) + | _ => (AssocMap.add (fst (mod_reset m)) (ZToValue 1 0) empty_assocmap) end. -(*Definition module_run (n : nat) (m : module) : res assocset := - mi_run (initial_fextclk m) empty_assocset (mod_body m) n.*) +(*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. @@ -601,18 +587,18 @@ Definition genv := Genv.t unit unit. Inductive step : state -> state -> Prop := | step_module : forall m stvar stvar' cycle f assoc0 assoc1 assoc2 nbassoc, - mis_stepp (f cycle) (mkassociations assoc0 empty_assocset) + mis_stepp (f cycle) (mkassociations assoc0 empty_assocmap) m.(mod_body) (mkassociations assoc1 nbassoc) -> - assoc2 = merge_assocset nbassoc assoc1 -> + assoc2 = merge_assocmap nbassoc assoc1 -> Some stvar' = assoc2!(fst m.(mod_state)) -> - step (State m assoc0 empty_assocset f cycle stvar) - (State m assoc2 empty_assocset f (S cycle) stvar') + step (State m assoc0 empty_assocmap f cycle stvar) + (State m assoc2 empty_assocmap f (S cycle) stvar') | step_finish : forall m assoc f cycle stvar result, assoc!(fst m.(mod_finish)) = Some (1'h"1") -> assoc!(fst m.(mod_return)) = Some result -> - step (State m assoc empty_assocset f cycle stvar) + step (State m assoc empty_assocmap f cycle stvar) (Finishstate result). Hint Constructors step : verilog. @@ -620,7 +606,7 @@ Hint Constructors step : verilog. | initial_state_intro: forall hmi tmi, hmi::tmi = mod_body m -> - initial_state m (State m hmi tmi empty_assocset empty_assocset (initial_fextclk m) O xH). + initial_state m (State m hmi tmi empty_assocmap empty_assocmap (initial_fextclk m) O xH). (** A final state is a [Returnstate] with an empty call stack. *) -- cgit From 7dbed3811e57479c78e4e82806d343d9285576c1 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 8 May 2020 23:34:36 +0100 Subject: Add lessdef for values --- src/verilog/Value.v | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index a96d67c..cde98d4 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -19,8 +19,8 @@ (* begin hide *) From bbv Require Import Word. From bbv Require HexNotation WordScope. -From Coq Require Import ZArith.ZArith. -From compcert Require Import lib.Integers. +From Coq Require Import ZArith.ZArith FSets.FMapPositive. +From compcert Require Import lib.Integers common.Values. (* end hide *) (** * Value @@ -48,7 +48,7 @@ Definition wordToValue : forall sz : nat, word sz -> value := mkvalue. Definition valueToWord : forall v : value, word (vsize v) := vword. -Definition valueToNat (v : value) : nat := +Definition valueToNat (v :value) : nat := wordToNat (vword v). Definition natToValue sz (n : nat) : value := @@ -284,3 +284,10 @@ Module HexNotationValue. Notation "sz ''h' a" := (NToValue sz (hex a)) (at level 50). End HexNotationValue. + +Inductive val_value_lessdef: val -> value -> Prop := +| val_value_lessdef_int: + forall i v', + Integers.Int.unsigned i = valueToZ v' -> + val_value_lessdef (Vint i) v' +| lessdef_undef: forall v, val_value_lessdef Vundef v. -- cgit From ebcba2714fdba588fb31c5a5307a8f21b39e0ac5 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 8 May 2020 23:35:06 +0100 Subject: Add simulation diagram --- src/translation/HTLgenproof.v | 58 +++++++++++++++++++++++++++++++++++++++---- 1 file changed, 53 insertions(+), 5 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 898fafd..7a4ec34 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -16,14 +16,62 @@ * along with this program. If not, see . *) -From coqup Require Import HTLgenspec. +From coqup Require Import HTLgenspec Value AssocMap. From coqup Require HTL Verilog. -From compcert Require RTL. +From compcert Require RTL Registers Globalenvs AST. -Parameter match_assocset : RTL.regset -> Verilog.assocset -> Prop. +Import AssocMapNotation. + +Inductive match_assocmaps : RTL.regset -> assocmap -> Prop := +| match_assocmap : forall r rs am, + val_value_lessdef (Registers.Regmap.get r rs) am#r -> + match_assocmaps rs am. Inductive match_states : RTL.state -> HTL.state -> Prop := | match_state : forall (rs : RTL.regset) assoc sf f sp rs mem m st, - match_assocset rs assoc -> + match_assocmaps rs assoc -> + tr_module f m -> match_states (RTL.State sf f sp st rs mem) - (HTL.State m st assoc). + (HTL.State m st assoc) +| match_returnstate : forall v v' stack m, + val_value_lessdef v v' -> + match_states (RTL.Returnstate stack v m) (HTL.Returnstate v'). + +Inductive match_program : RTL.program -> HTL.module -> Prop := + match_program_intro : + forall ge p b m f, + ge = Globalenvs.Genv.globalenv p -> + Globalenvs.Genv.find_symbol ge p.(AST.prog_main) = Some b -> + Globalenvs.Genv.find_funct_ptr ge b = Some (AST.Internal f) -> + tr_module f m -> + match_program p m. + +Section CORRECTNESS. + + Variable prog : RTL.program. + Variable tprog : HTL.module. + + Hypothesis TRANSL : match_program prog tprog. + + Let ge : RTL.genv := Globalenvs.Genv.globalenv prog. + Let tge : HTL.genv := HTL.genv_empty. + + (** The proof of semantic preservation for the translation of instructions + is a simulation argument based on diagrams of the following form: +<< + I /\ P + code st rs ---------------- State m st assoc + || | + || | + || | + \/ v + code st rs' --------------- State m st assoc' + I /\ Q +>> + where [tr_code c data control fin rtrn st] is assumed to hold. + *) + + Definition transl_code_prop : Prop := + HTL.step tge (HTL.State) + +End CORRECTNESS. -- cgit From b688e8bbf82e694cf62c11a117b67706d39f8459 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 20 May 2020 10:50:20 +0100 Subject: Update coq version to 11 --- default.nix | 4 ++-- lib/CompCert | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/default.nix b/default.nix index 961bdec..cd37308 100644 --- a/default.nix +++ b/default.nix @@ -1,8 +1,8 @@ with import {}; let - ncoq = coq_8_10; - ncoqPackages = coqPackages_8_10; + ncoq = coq_8_11; + ncoqPackages = coqPackages_8_11; bbv = ncoqPackages.callPackage ( { coq, stdenv, fetchFromGitHub }: stdenv.mkDerivation { diff --git a/lib/CompCert b/lib/CompCert index e134692..8f8250d 160000 --- a/lib/CompCert +++ b/lib/CompCert @@ -1 +1 @@ -Subproject commit e13469254ced9e4980ee54769d5b8b8729efc850 +Subproject commit 8f8250d076c9bcb15954c978dd1c351dd948326c -- cgit From 333b306421d204e69deb1308352e31ffe53d2287 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 20 May 2020 10:50:37 +0100 Subject: Add theorems about merge --- src/verilog/AssocMap.v | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index e550072..bd61c8e 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -30,14 +30,24 @@ Module AssocMapExt. Variable elt : Type. - Definition merge : t elt -> t elt -> t elt := fold (@add elt). - Definition find_default (a : elt) (k : reg) (m : t elt) : elt := match find k m with | None => a | Some v => v end. + Definition merge : t elt -> t elt -> t elt := fold (@add elt). + + Lemma merge_add_assoc : + forall am am' r v x, + find x (merge (add r v am) am') = find x (add r v (merge am am')). + Admitted. + + Lemma merge_base : + forall am, + merge (empty elt) am = am. + Proof. intros. unfold merge. apply fold_1. Qed. + End Operations. End AssocMapExt. -- cgit From 75c4a7bcad1ffd3ba69a3b514fb71590c7a523c9 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 20 May 2020 10:50:58 +0100 Subject: Fix definitions in Value and add lemmas --- src/verilog/Value.v | 42 +++++++++++++++++++++++++++++++++++------- 1 file changed, 35 insertions(+), 7 deletions(-) diff --git a/src/verilog/Value.v b/src/verilog/Value.v index cde98d4..0ce0bc5 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -70,17 +70,14 @@ Definition uvalueToZ (v : value) : Z := uwordToZ (vword v). Definition posToValue sz (p : positive) : value := - mkvalue sz (posToWord sz p). + ZToValue sz (Zpos p). Definition posToValueAuto (p : positive) : value := - let size := Z.to_nat (Z.succ (log_inf p)) in - mkvalue size (Word.posToWord size p). + let size := Pos.to_nat (Pos.size p) in + ZToValue size (Zpos p). Definition valueToPos (v : value) : positive := - match valueToZ v with - | Zpos p => p - | _ => 1 - end. + Z.to_pos (uvalueToZ v). Definition intToValue (i : Integers.int) : value := ZToValue Int.wordsize (Int.unsigned i). @@ -88,6 +85,12 @@ Definition intToValue (i : Integers.int) : value := Definition valueToInt (i : value) : Integers.int := Int.repr (valueToZ i). +Definition valToValue (v : Values.val) := + match v with + | Values.Vint i => intToValue i + | _ => ZToValue 32 0%Z + 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 @@ -291,3 +294,28 @@ Inductive val_value_lessdef: val -> value -> Prop := Integers.Int.unsigned i = valueToZ v' -> val_value_lessdef (Vint i) v' | lessdef_undef: forall v, val_value_lessdef Vundef v. + +Search Z positive. + +Search "wordToZ". + +Lemma uvalueToZ_ZToValue : + forall n z, + (0 <= z < 2 ^ Z.of_nat n)%Z -> + uvalueToZ (ZToValue n z) = z. +Proof. + intros. unfold uvalueToZ, ZToValue. simpl. + apply uwordToZ_ZToWord. apply H. +Qed. + +Lemma valueToPos_posToValueAuto : + forall p, valueToPos (posToValueAuto p) = p. +Proof. + intros. unfold valueToPos, posToValueAuto. + rewrite uvalueToZ_ZToValue. auto. rewrite positive_nat_Z. + split. apply Zle_0_pos. + + assert (p < 2 ^ (Pos.size p))%positive. apply Pos.size_gt. + inversion H. rewrite <- Z.compare_lt_iff. rewrite <- H1. + simpl. rewrite <- Pos2Z.inj_pow_pos. trivial. +Qed. -- cgit From 43f766ea8aff8309d94173cc1e2670eb8ddce68f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 20 May 2020 10:51:14 +0100 Subject: Switch position of empty rule --- src/verilog/Verilog.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index c0bce25..c1e0a79 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -244,14 +244,14 @@ Inductive expr_runp : fext -> assocmap -> expr -> value -> Prop := | erun_Vlit : forall fext assoc v, expr_runp fext assoc (Vlit v) v - | erun_Vvar_empty : - forall fext assoc r sz, - assoc!r = None -> - expr_runp fext assoc (Vvar r) (ZToValue sz 0) | erun_Vvar : forall fext assoc v r, assoc!r = Some v -> expr_runp fext assoc (Vvar r) v + | erun_Vvar_empty : + forall fext assoc r sz, + assoc!r = None -> + expr_runp fext assoc (Vvar r) (ZToValue sz 0) | erun_Vinputvar : forall fext assoc r v, fext!r = Some v -> -- cgit From d76ea8e7a52f01a1be2f5b2ec1c5367a42aa0f65 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 20 May 2020 10:51:33 +0100 Subject: Fix the semantics to properly evaluate the state --- src/verilog/HTL.v | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v index 147d24b..36e4434 100644 --- a/src/verilog/HTL.v +++ b/src/verilog/HTL.v @@ -63,7 +63,7 @@ Inductive state : Type := Inductive step : genv -> state -> Events.trace -> state -> Prop := | step_module : - forall g t m st ctrl data assoc0 assoc1 assoc2 assoc3 nbassoc0 nbassoc1 f, + forall g t m st ctrl data assoc0 assoc1 assoc2 assoc3 nbassoc0 nbassoc1 f stval pstval, m.(mod_controllogic)!st = Some ctrl -> m.(mod_datapath)!st = Some data -> Verilog.stmnt_runp f @@ -75,7 +75,9 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop := data (Verilog.mkassociations assoc2 nbassoc1) -> assoc3 = merge_assocmap nbassoc1 assoc2 -> - step g (State m st assoc0) t (State m st assoc3) + assoc3!(m.(mod_st)) = Some stval -> + valueToPos stval = pstval -> + step g (State m st assoc0) t (State m pstval assoc3) | step_finish : forall g t m st assoc retval, assoc!(m.(mod_finish)) = Some (1'h"1") -> -- cgit From 60ab47c6ff0db1b690afe6fd2c2b63cf5843ced1 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 20 May 2020 10:51:49 +0100 Subject: Add proof of translation correctness Modified-by: Yann Herklotz --- src/translation/HTLgenproof.v | 201 ++++++++++++++++++++++++++++++++++++++++-- src/translation/HTLgenspec.v | 16 ++-- 2 files changed, 200 insertions(+), 17 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 7a4ec34..ff7c24d 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -23,14 +23,22 @@ From compcert Require RTL Registers Globalenvs AST. Import AssocMapNotation. Inductive match_assocmaps : RTL.regset -> assocmap -> Prop := -| match_assocmap : forall r rs am, - val_value_lessdef (Registers.Regmap.get r rs) am#r -> +| match_assocmap : forall rs am, + (forall r, val_value_lessdef (Registers.Regmap.get r rs) am#r) -> match_assocmaps rs am. +Definition state_st_wf (m : HTL.module) (s : HTL.state) := + forall st assoc, + s = HTL.State m st assoc -> + assoc!(m.(HTL.mod_st)) = Some (posToValue 32 st). + Inductive match_states : RTL.state -> HTL.state -> Prop := -| match_state : forall (rs : RTL.regset) assoc sf f sp rs mem m st, - match_assocmaps rs assoc -> - tr_module f m -> +| match_state : forall (rs : RTL.regset) assoc sf f sp rs mem m st + (MASSOC : match_assocmaps rs assoc) + (TF : tr_module f m) + (TC : tr_code f.(RTL.fn_code) st m.(HTL.mod_datapath) m.(HTL.mod_controllogic) + m.(HTL.mod_finish) m.(HTL.mod_return) m.(HTL.mod_st)) + (WF : state_st_wf m (HTL.State m st assoc)), match_states (RTL.State sf f sp st rs mem) (HTL.State m st assoc) | match_returnstate : forall v v' stack m, @@ -46,6 +54,80 @@ Inductive match_program : RTL.program -> HTL.module -> Prop := tr_module f m -> match_program p m. +Lemma regs_lessdef_regs : + forall rs1 rs2 n v, + match_assocmaps rs1 rs2 -> + match_assocmaps rs1 (AssocMap.add n v rs2). +Admitted. + +Lemma regs_add_match : + forall rs am r v v', + match_assocmaps rs am -> + val_value_lessdef v v' -> + match_assocmaps (Registers.Regmap.set r v rs) (AssocMap.add r v' am). +Admitted. + +Lemma merge_inj : + forall am am' r v, + merge_assocmap (AssocMap.add r v am) am' = AssocMap.add r v (merge_assocmap am am'). +Admitted. + +Lemma valToValue_lessdef : + forall v, + val_value_lessdef v (valToValue v). +Admitted. + +Lemma assocmap_merge_add : + forall k v assoc, + AssocMap.add k v assoc = merge_assocmap (AssocMap.add k v empty_assocmap) assoc. +Admitted. + +(* Need to eventually move posToValue 32 to posToValueAuto, as that has this proof. *) +Lemma assumption_32bit : + forall v, + valueToPos (posToValue 32 v) = v. +Admitted. + +Lemma regset_assocmap_get : + forall r rs am v v', + match_assocmaps rs am -> + v = Registers.Regmap.get r rs -> + v' = am#r -> + val_value_lessdef v v'. +Proof. inversion 1. intros. subst. apply H0. Qed. + +Lemma regset_assocmap_wf : + forall r rs am i, + match_assocmaps rs am -> + Values.Vint i <> Registers.Regmap.get r rs -> + am!r = None. +Admitted. + +Lemma regset_assocmap_wf2 : + forall r rs am i, + match_assocmaps rs am -> + Values.Vint i = Registers.Regmap.get r rs -> + am!r = Some (intToValue i). +Admitted. + +Lemma access_merge_assocmap : + forall nb r v am, + nb!r = Some v -> + (merge_assocmap nb am) ! r = Some v. +Admitted. + +Lemma st_greater_than_res : + forall m res, + (HTL.mod_st m) <> res. +Admitted. + +Ltac subst_eq_hyp := + match goal with + H1: ?x = Some ?i, H2: ?x = Some ?j |- _ => + let H := fresh "H" in + rewrite H1 in H2; injection H2; intro H; clear H2; subst + end. + Section CORRECTNESS. Variable prog : RTL.program. @@ -56,22 +138,123 @@ Section CORRECTNESS. Let ge : RTL.genv := Globalenvs.Genv.globalenv prog. Let tge : HTL.genv := HTL.genv_empty. + Lemma eval_correct : + forall sp op rs args m v e assoc f, + Op.eval_operation ge sp op +(List.map (fun r : BinNums.positive => Registers.Regmap.get r rs) args) m = Some v -> + tr_op op args e -> + Verilog.expr_runp f assoc e (valToValue v). + Admitted. + (** The proof of semantic preservation for the translation of instructions is a simulation argument based on diagrams of the following form: << - I /\ P + match_states code st rs ---------------- State m st assoc || | || | || | \/ v code st rs' --------------- State m st assoc' - I /\ Q + I >> 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_code_prop : Prop := - HTL.step tge (HTL.State) + Definition transl_instr_prop (instr : RTL.instruction) : Prop := + forall m assoc fin rtrn st stmt trans, + tr_instr fin rtrn st instr stmt trans -> + exists assoc', + HTL.step tge (HTL.State m st assoc) Events.E0 (HTL.State m st assoc'). + + Theorem transl_step_correct: + forall (S1 : RTL.state) t S2, + RTL.step ge S1 t S2 -> + forall (R1 : HTL.state), + match_states S1 R1 -> + exists R2, HTL.step tge R1 t R2 /\ match_states S2 R2. + Proof. + induction 1; intros R1 MSTATE. + - (* Inop *) + inversion MSTATE. subst. + econstructor. + split. + + inversion TC. + eapply HTL.step_module; subst_eq_hyp; remember (HTL.mod_st m0) as st; inversion H3; subst. + apply H2. + apply H1. + (* processing of state *) + econstructor. + simpl. trivial. + econstructor. trivial. + inversion H3. subst. + econstructor. + + (* prove merge *) + apply assocmap_merge_add. + + (* prove stval *) + apply AssocMap.gss. + apply assumption_32bit. + + (* prove match_state *) + constructor. + apply rs. + apply regs_lessdef_regs. + assumption. + inversion TF. simpl. apply H0. + assumption. + unfold state_st_wf. intros. inversion H0. subst. + apply AssocMap.gss. + - (* Iop *) + inversion MSTATE. subst. + econstructor. + split. + + inversion TC. + eapply HTL.step_module; subst_eq_hyp; remember (HTL.mod_st m0) as st; inversion H4; subst. + apply H3. + apply H2. + econstructor. + simpl. trivial. + constructor. trivial. + econstructor. + simpl. trivial. + eapply eval_correct. apply H0. apply H10. trivial. trivial. + apply access_merge_assocmap. rewrite AssocMap.gso. + apply AssocMap.gss. + apply st_greater_than_res. + trivial. + + (* match_states *) + assert (pc' = valueToPos (posToValue 32 pc')). symmetry. apply assumption_32bit. + rewrite <- H1. + constructor. apply rs0. + rewrite merge_inj. + apply regs_add_match. + rewrite merge_inj. + unfold merge_assocmap. rewrite AssocMapExt.merge_base. + apply regs_lessdef_regs. + assumption. + apply valToValue_lessdef. + + inversion TF. simpl. apply H2. + assumption. + + unfold state_st_wf. intros. inversion H2. subst. + rewrite merge_inj. + rewrite AssocMap.gso. + rewrite merge_inj. + apply AssocMap.gss. + apply st_greater_than_res. + - inversion MSTATE. inversion TC. subst. + econstructor. constructor. + inversion H12; subst_eq_hyp; discriminate. + apply match_state. apply rs0. + apply regs_add_match. apply MASSOC. apply valToValue_lessdef. + inversion TF. rewrite H3. End CORRECTNESS. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index fb1c70d..e2d788c 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -70,20 +70,20 @@ Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Pr tr_instr fin rtrn st (RTL.Ireturn (Some r)) (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r))) Vskip. -Inductive tr_code (c : RTL.code) (stmnts trans : PositiveMap.t stmnt) +Inductive tr_code (c : RTL.code) (pc : RTL.node) (stmnts trans : PositiveMap.t stmnt) (fin rtrn st : reg) : Prop := tr_code_intro : - forall i s t n, - Maps.PTree.get n c = Some i -> - stmnts!n = Some s -> - trans!n = Some t -> + forall i s t, + Maps.PTree.get pc c = Some i -> + stmnts!pc = Some s -> + trans!pc = Some t -> tr_instr fin rtrn st i s t -> - tr_code c stmnts trans fin rtrn st. + tr_code c pc stmnts trans fin rtrn st. Inductive tr_module (f : RTL.function) : module -> Prop := tr_module_intro : - forall c data control fin rtrn st, - tr_code c data control fin rtrn st -> + forall data control fin rtrn st, + (forall pc, tr_code f.(RTL.fn_code) pc data control fin rtrn st) -> tr_module f (mkmodule f.(RTL.fn_params) data -- cgit From 39a4348657c2f3efb3feafe9cf65b0f2a1a263c2 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 21 May 2020 23:10:45 +0100 Subject: Finish the proof with most assumptions --- src/translation/HTLgenproof.v | 177 +++++++++++++++++++++++++++++++++++------- src/translation/HTLgenspec.v | 2 +- src/verilog/AssocMap.v | 17 ++-- 3 files changed, 161 insertions(+), 35 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index ff7c24d..cb621a8 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -88,6 +88,11 @@ Lemma assumption_32bit : valueToPos (posToValue 32 v) = v. Admitted. +Lemma assumption_32bit_bool : + forall b, + valueToBool (boolToValue 32 b) = b. +Admitted. + Lemma regset_assocmap_get : forall r rs am v v', match_assocmaps rs am -> @@ -110,15 +115,19 @@ Lemma regset_assocmap_wf2 : am!r = Some (intToValue i). Admitted. -Lemma access_merge_assocmap : - forall nb r v am, - nb!r = Some v -> - (merge_assocmap nb am) ! r = Some v. -Admitted. - Lemma st_greater_than_res : forall m res, - (HTL.mod_st m) <> res. + HTL.mod_st m <> res. +Admitted. + +Lemma finish_not_return : + forall m, + HTL.mod_finish m <> HTL.mod_return m. +Admitted. + +Lemma finish_not_res : + forall m r, + HTL.mod_finish m <> r. Admitted. Ltac subst_eq_hyp := @@ -128,6 +137,9 @@ Ltac subst_eq_hyp := rewrite H1 in H2; injection H2; intro H; clear H2; subst end. +Ltac unfold_merge := + try (repeat (rewrite merge_inj)); unfold merge_assocmap; rewrite AssocMapExt.merge_base. + Section CORRECTNESS. Variable prog : RTL.program. @@ -146,6 +158,13 @@ Section CORRECTNESS. Verilog.expr_runp f assoc e (valToValue v). Admitted. + Lemma eval_cond_correct : + forall cond (args : list Registers.reg) s1 c s' i rs args m b f assoc, + Veriloggen.translate_condition cond args s1 = Veriloggen.OK c s' i -> + Op.eval_condition cond (List.map (fun r : BinNums.positive => Registers.Regmap.get r rs) args) m = Some b -> + Verilog.expr_runp f assoc c (boolToValue 32 b). + Admitted. + (** The proof of semantic preservation for the translation of instructions is a simulation argument based on diagrams of the following form: << @@ -156,7 +175,7 @@ Section CORRECTNESS. || | \/ v code st rs' --------------- State m st assoc' - I + match_states >> where [tr_code c data control fin rtrn st] is assumed to hold. @@ -174,16 +193,17 @@ Section CORRECTNESS. RTL.step ge S1 t S2 -> forall (R1 : HTL.state), match_states S1 R1 -> - exists R2, HTL.step tge R1 t R2 /\ match_states S2 R2. + exists R2, Smallstep.plus HTL.step tge R1 t R2 /\ match_states S2 R2. Proof. - induction 1; intros R1 MSTATE. + induction 1; intros R1 MSTATE; + try (inversion MSTATE; inversion TC; inversion H12; subst_eq_hyp; discriminate). - (* Inop *) inversion MSTATE. subst. econstructor. split. - + apply Smallstep.plus_one. inversion TC. - eapply HTL.step_module; subst_eq_hyp; remember (HTL.mod_st m0) as st; inversion H3; subst. + eapply HTL.step_module; subst_eq_hyp; inversion H3; subst. apply H2. apply H1. (* processing of state *) @@ -205,17 +225,17 @@ Section CORRECTNESS. apply rs. apply regs_lessdef_regs. assumption. - inversion TF. simpl. apply H0. assumption. + inversion TF. simpl. apply H0. unfold state_st_wf. intros. inversion H0. subst. apply AssocMap.gss. - (* Iop *) inversion MSTATE. subst. econstructor. split. - + apply Smallstep.plus_one. inversion TC. - eapply HTL.step_module; subst_eq_hyp; remember (HTL.mod_st m0) as st; inversion H4; subst. + eapply HTL.step_module; subst_eq_hyp; inversion H4; subst. apply H3. apply H2. econstructor. @@ -224,7 +244,8 @@ Section CORRECTNESS. econstructor. simpl. trivial. eapply eval_correct. apply H0. apply H10. trivial. trivial. - apply access_merge_assocmap. rewrite AssocMap.gso. + unfold_merge. + rewrite AssocMap.gso. apply AssocMap.gss. apply st_greater_than_res. trivial. @@ -233,28 +254,128 @@ Section CORRECTNESS. assert (pc' = valueToPos (posToValue 32 pc')). symmetry. apply assumption_32bit. rewrite <- H1. constructor. apply rs0. - rewrite merge_inj. + unfold_merge. apply regs_add_match. - rewrite merge_inj. - unfold merge_assocmap. rewrite AssocMapExt.merge_base. apply regs_lessdef_regs. assumption. apply valToValue_lessdef. - - inversion TF. simpl. apply H2. assumption. + inversion TF. simpl. apply H2. unfold state_st_wf. intros. inversion H2. subst. - rewrite merge_inj. + unfold_merge. rewrite AssocMap.gso. - rewrite merge_inj. apply AssocMap.gss. apply st_greater_than_res. - - inversion MSTATE. inversion TC. subst. - econstructor. constructor. - inversion H12; subst_eq_hyp; discriminate. + - inversion MSTATE; inversion TC; + inversion H12; subst_eq_hyp; inversion H13. + - inversion MSTATE. inversion TC. subst_eq_hyp. inversion H12. subst. + econstructor. split. apply Smallstep.plus_one. + eapply HTL.step_module. subst. apply H11. apply H10. + eapply Verilog.stmnt_runp_Vnonblock with + (rhsval := if b then posToValue 32 ifso else posToValue 32 ifnot). + simpl. trivial. + destruct b. + eapply Verilog.erun_Vternary_true. + eapply eval_cond_correct. apply args. + apply H7. apply H0. + constructor. + apply assumption_32bit_bool. + eapply Verilog.erun_Vternary_false. + eapply eval_cond_correct. apply args. + apply H7. apply H0. + constructor. + apply assumption_32bit_bool. + trivial. + constructor. + trivial. + unfold_merge. + apply AssocMap.gss. + trivial. + + destruct b. + rewrite assumption_32bit. + apply match_state. apply rs0. + unfold_merge. + apply regs_lessdef_regs. assumption. + assumption. + + inversion TF. simpl. apply H1. + + unfold state_st_wf. intros. inversion H1. + subst. unfold_merge. + apply AssocMap.gss. + rewrite assumption_32bit. apply match_state. apply rs0. - apply regs_add_match. apply MASSOC. apply valToValue_lessdef. - inversion TF. rewrite H3. + unfold_merge. + apply regs_lessdef_regs. assumption. + assumption. + + inversion TF. simpl. apply H1. + + unfold state_st_wf. intros. inversion H1. + subst. unfold_merge. + apply AssocMap.gss. + + - (* Return *) + inversion MSTATE. inversion TC. subst_eq_hyp. + inversion H11. subst. + econstructor. split. + eapply Smallstep.plus_two. + + eapply HTL.step_module. + apply H10. + apply H9. + constructor. + econstructor; simpl; trivial. + econstructor; simpl; trivial. + constructor. + econstructor; simpl; trivial. + constructor. + unfold_merge. + trivial. + rewrite AssocMap.gso. + rewrite AssocMap.gso. + unfold state_st_wf in WF. apply WF. trivial. + apply st_greater_than_res. apply st_greater_than_res. trivial. + + apply HTL.step_finish. + rewrite AssocMap.gso. + apply AssocMap.gss. + apply finish_not_return. + apply AssocMap.gss. + rewrite Events.E0_left. trivial. + + constructor. constructor. + destruct (assoc!r) eqn:?. + inversion H11. subst. + econstructor. split. + eapply Smallstep.plus_two. + eapply HTL.step_module. + apply H10. + apply H9. + constructor. + econstructor; simpl; trivial. + econstructor; simpl; trivial. + constructor. + econstructor; simpl; trivial. + apply Verilog.erun_Vvar. + rewrite AssocMap.gso. + apply Heqo. apply not_eq_sym. apply finish_not_res. + unfold_merge. + trivial. + rewrite AssocMap.gso. + rewrite AssocMap.gso. + unfold state_st_wf in WF. apply WF. trivial. + apply st_greater_than_res. apply st_greater_than_res. trivial. + + apply HTL.step_finish. + rewrite AssocMap.gso. + apply AssocMap.gss. + apply finish_not_return. + apply AssocMap.gss. + rewrite Events.E0_left. trivial. + + constructor. simpl. Search Registers.Regmap.get. End CORRECTNESS. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index e2d788c..a7d65fc 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -64,7 +64,7 @@ Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Pr translate_condition cond args s = OK c s' i -> tr_instr fin rtrn st (RTL.Icond cond args n1 n2) Vskip (state_cond st c n1 n2) | tr_instr_Ireturn_None : - tr_instr fin rtrn st (RTL.Ireturn None) (block fin (Vlit (ZToValue 1%nat 1%Z))) Vskip + tr_instr fin rtrn st (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vlit (ZToValue 1%nat 0%Z)))) Vskip | tr_instr_Ireturn_Some : forall r, tr_instr fin rtrn st (RTL.Ireturn (Some r)) diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index bd61c8e..9caa2d1 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -36,12 +36,17 @@ Module AssocMapExt. | Some v => v end. - Definition merge : t elt -> t elt -> t elt := fold (@add elt). - - Lemma merge_add_assoc : - forall am am' r v x, - find x (merge (add r v am) am') = find x (add r v (merge am am')). - Admitted. + Definition merge (m1 m2 : t elt) : t elt := + fold_right (fun el m => match el with (k, v) => add k v m end) m2 (elements m1). + + Lemma merge_add_assoc2 : + forall am am' r v, + merge (add r v am) am' = add r v (merge am am'). + Proof. + intros. + unfold merge. + Search fold_right. + apply SetoidList.fold_right_add2. Lemma merge_base : forall am, -- cgit From d1bea5e0f88ad2cfcd0b0667fb11279100d02e6b Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 24 May 2020 13:36:05 +0100 Subject: Add statemonad declaration --- src/common/Monad.v | 42 ++++++++++++++++++++++++++++++++++ src/common/Statemonad.v | 61 +++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 103 insertions(+) create mode 100644 src/common/Monad.v create mode 100644 src/common/Statemonad.v diff --git a/src/common/Monad.v b/src/common/Monad.v new file mode 100644 index 0000000..773918a --- /dev/null +++ b/src/common/Monad.v @@ -0,0 +1,42 @@ +From Coq Require Import Lists.List. + +Module Type Monad. + + Parameter mon : Type -> Type. + + Parameter ret : forall (A : Type) (x : A), mon A. + Arguments ret [A]. + + Parameter bind : forall (A B : Type) (f : mon A) (g : A -> mon B), mon B. + Arguments bind [A B]. + + Parameter bind2 : forall (A B C: Type) (f: mon (A * B)) (g: A -> B -> mon C), mon C. + Arguments bind2 [A B C]. + +End Monad. + +Module MonadExtra(M : Monad). + Import M. + + Module MonadNotation. + + Notation "'do' X <- A ; B" := + (bind A (fun X => B)) + (at level 200, X ident, A at level 100, B at level 200). + Notation "'do' ( X , Y ) <- A ; B" := + (bind2 A (fun X Y => B)) + (at level 200, X ident, Y ident, A at level 100, B at level 200). + + End MonadNotation. + Import MonadNotation. + + Fixpoint traverselist {A B: Type} (f: A -> mon B) (l: list A) {struct l}: mon (list B) := + match l with + | nil => ret nil + | x::xs => + do r <- f x; + do rs <- traverselist f xs; + ret (r::rs) + end. + +End MonadExtra. diff --git a/src/common/Statemonad.v b/src/common/Statemonad.v new file mode 100644 index 0000000..ed1b9e7 --- /dev/null +++ b/src/common/Statemonad.v @@ -0,0 +1,61 @@ +From compcert Require Errors. +From coqup Require Import Monad. +From Coq Require Import Lists.List. + +Module Type State. + Parameter st : Type. + Parameter st_prop : st -> st -> Prop. + + Axiom st_refl : forall s, st_prop s s. + Axiom st_trans : forall s1 s2 s3, st_prop s1 s2 -> st_prop s2 s3 -> st_prop s1 s3. +End State. + +Module Statemonad(S : State) <: Monad. + + Inductive res (A: Type) (s: S.st): Type := + | Error : Errors.errmsg -> res A s + | OK : A -> forall (s' : S.st), S.st_prop s s' -> res A s. + + Arguments OK [A s]. + Arguments Error [A s]. + + Definition mon (A: Type) : Type := forall (s: S.st), res A s. + + Definition ret {A: Type} (x: A) : mon A := + fun (s : S.st) => OK x s (S.st_refl s). + + Definition bind {A B: Type} (f: mon A) (g: A -> mon B) : mon B := + fun (s : S.st) => + match f s with + | Error msg => Error msg + | OK a s' i => + match g a s' with + | Error msg => Error msg + | OK b s'' i' => OK b s'' (S.st_trans s s' s'' i i') + end + end. + + Definition bind2 {A B C: Type} (f: mon (A * B)) (g: A -> B -> mon C) : mon C := + bind f (fun xy => g (fst xy) (snd xy)). + + Definition handle_error {A: Type} (f g: mon A) : mon A := + fun (s : S.st) => + match f s with + | OK a s' i => OK a s' i + | Error _ => g s + end. + + Definition error {A: Type} (err: Errors.errmsg) : mon A := fun (s: S.st) => Error err. + + Definition get : mon S.st := fun s => OK s s (S.st_refl s). + + Definition set (s: S.st) (i: forall s', S.st_prop s' s) : mon unit := + fun s' => OK tt s (i s'). + + Definition run_mon {A: Type} (s: S.st) (m: mon A): Errors.res A := + match m s with + | OK a s' i => Errors.OK a + | Error err => Errors.Error err + end. + +End Statemonad. -- cgit From bb8a935f9143e65102a3a498a96dd13a3b8a4801 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 24 May 2020 13:36:53 +0100 Subject: Add HTLgen --- src/translation/HTLgen.v | 338 ++++++++++++++++++++++++++++++++++++++++++++++- src/verilog/AssocMap.v | 9 +- 2 files changed, 341 insertions(+), 6 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 8eac48c..7bb5030 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -16,8 +16,344 @@ * along with this program. If not, see . *) -From coqup Require Import HTL Coquplib. From compcert Require Import Maps. From compcert Require Errors. From compcert Require Import AST RTL. +From coqup Require Import Verilog HTL Coquplib AssocMap Value Statemonad. +Hint Resolve AssocMap.gempty : htlh. +Hint Resolve AssocMap.gso : htlh. +Hint Resolve AssocMap.gss : htlh. +Hint Resolve Ple_refl : htlh. +Hint Resolve Ple_succ : htlh. + +Record state: Type := mkstate { + st_st : reg; + st_freshreg: reg; + st_freshstate: node; + st_datapath: datapath; + st_controllogic: controllogic; +}. + +Definition init_state (st : reg) : state := + mkstate st + 1%positive + 1%positive + (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_datapath)!n = None \/ s2.(st_datapath)!n = s1.(st_datapath)!n) -> + (forall n, + s1.(st_controllogic)!n = None + \/ s2.(st_controllogic)!n = s1.(st_controllogic)!n) -> + st_incr s1 s2. + Hint Constructors st_incr : htlh. + + Definition st_prop := st_incr. + 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. + - rewrite H1. rewrite H. reflexivity. + - intros. destruct H4 with n. + + left; assumption. + + destruct H8 with n; + [ rewrite <- H0; left; assumption + | right; rewrite <- H0; apply H10 + ]. + - intros. destruct H5 with n. + + left; assumption. + + destruct H9 with n. + * rewrite <- H0. left. assumption. + * right. rewrite <- H0. apply H10. + Qed. + +End HTLState. +Import HTLState. + +Module HTLMonad := Statemonad(HTLState). +Import HTLMonad. + +Module HTLMonadExtra := Monad.MonadExtra(HTLMonad). +Import HTLMonadExtra. +Import MonadNotation. + +Definition state_goto (st : reg) (n : node) : stmnt := + Vnonblock (Vvar st) (Vlit (posToValue 32 n)). + +Definition state_cond (st : reg) (c : expr) (n1 n2 : node) : stmnt := + Vnonblock (Vvar st) (Vternary c (posToExpr 32 n1) (posToExpr 32 n2)). + +Definition check_empty_node_datapath: + forall (s: state) (n: node), { s.(st_datapath)!n = None } + { True }. +Proof. + intros. case (s.(st_datapath)!n); intros. right; auto. left; auto. +Defined. + +Definition check_empty_node_controllogic: + forall (s: state) (n: node), { s.(st_controllogic)!n = None } + { True }. +Proof. + intros. case (s.(st_controllogic)!n); intros. right; auto. left; auto. +Defined. + +Lemma add_instr_state_incr : + forall s n n' st, + (st_datapath s)!n = None -> + (st_controllogic s)!n = None -> + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + (AssocMap.add n st s.(st_datapath)) + (AssocMap.add 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_datapath s n, check_empty_node_controllogic s n with + | left STM, left TRANS => + OK tt (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + (AssocMap.add n st s.(st_datapath)) + (AssocMap.add n (state_goto s.(st_st) n') s.(st_controllogic))) + (add_instr_state_incr s n n' st STM TRANS) + | _, _ => Error (Errors.msg "HTL.add_instr") + 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 32%nat 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 "Veriloggen: 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 "Veriloggen: 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_comparison c args + | Op.Ccompimm c i, _ => translate_comparison_imm c args i + | Op.Ccompuimm c i, _ => translate_comparison_imm c args i + | Op.Cmaskzero n, _ => error (Errors.msg "Veriloggen: condition instruction not implemented: Cmaskzero") + | Op.Cmasknotzero n, _ => error (Errors.msg "Veriloggen: condition instruction not implemented: Cmasknotzero") + | _, _ => error (Errors.msg "Veriloggen: condition instruction not implemented: other") + end. + +Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon expr := + match a, args with + | Op.Aindexed off, r1::nil => ret (boplitz Vadd r1 off) + | Op.Aindexed2 off, r1::r2::nil => ret (Vbinop Vadd (Vvar r1) (boplitz Vadd r2 off)) + | Op.Ascaled scale offset, r1::nil => + ret (Vbinop Vadd (boplitz Vadd r1 scale) (Vlit (ZToValue 32%nat offset))) + | Op.Aindexed2scaled scale offset, r1::r2::nil => + ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) + | _, _ => error (Errors.msg "Veriloggen: eff_addressing instruction not implemented: other") + end. + +(** Translate an instruction to a statement. *) +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, _ => error (Errors.msg "Veriloggen: Instruction not implemented: Omulhs") + | Op.Omulhu, _ => error (Errors.msg "Veriloggen: Instruction not implemented: Omulhu") + | 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 => error (Errors.msg "Veriloggen: Instruction not implemented: Oshrximm") + | Op.Oshru, r1::r2::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Oshru") + | Op.Oshruimm n, r::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Oshruimm") + | Op.Ororimm n, r::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Ororimm") + | Op.Oshldimm n, r::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Oshldimm") + | Op.Ocmp c, _ => translate_condition c args + | Op.Olea a, _ => translate_eff_addressing a args + | _, _ => error (Errors.msg "Veriloggen: Instruction not implemented: other") + end. + +Lemma add_branch_instr_state_incr: + forall s e n n1 n2, + (st_datapath s) ! n = None -> + (st_controllogic s) ! n = None -> + st_incr s (mkstate + s.(st_st) + (st_freshreg s) + (st_freshstate s) + (AssocMap.add n Vskip (st_datapath s)) + (AssocMap.add 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_datapath s n, check_empty_node_controllogic s n with + | left NSTM, left NTRANS => + OK tt (mkstate + s.(st_st) + (st_freshreg s) + (st_freshstate s) + (AssocMap.add n Vskip (st_datapath s)) + (AssocMap.add n (state_cond s.(st_st) e n1 n2) (st_controllogic s))) + (add_branch_instr_state_incr s e n n1 n2 NSTM NTRANS) + | _, _ => Error (Errors.msg "Veriloggen: add_branch_instr") + end. + +Definition transf_instr (fin rtrn: reg) (ni: node * instruction) : mon unit := + match ni with + (n, i) => + match i with + | Inop n' => add_instr n n' Vskip + | Iop op args dst n' => + do instr <- translate_instr op args; + add_instr n n' (block dst instr) + | Iload _ _ _ _ _ => error (Errors.msg "Loads are not implemented.") + | Istore _ _ _ _ _ => error (Errors.msg "Stores are not implemented.") + | Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") + | Itailcall _ _ _ => error (Errors.msg "Tailcalls are not implemented.") + | Ibuiltin _ _ _ _ => error (Errors.msg "Builtin functions not implemented.") + | Icond cond args n1 n2 => + do e <- translate_condition cond args; + add_branch_instr e n n1 n2 + | Ijumptable _ _ => error (Errors.msg "Jumptable not implemented") + | Ireturn r => + match r with + | Some r' => + add_instr n n (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r'))) + | None => + add_instr n n (block fin (Vlit (ZToValue 1%nat 1%Z))) + end + end + end. + +Lemma create_reg_state_incr: + forall s, + st_incr s (mkstate + s.(st_st) + (Pos.succ (st_freshreg s)) + (st_freshstate s) + (st_datapath s) + (st_controllogic s)). +Proof. constructor; simpl; auto with htlh. Qed. + +Definition create_reg: mon reg := + fun s => let r := s.(st_freshreg) in + OK r (mkstate + s.(st_st) + (Pos.succ r) + (st_freshstate s) + (st_datapath s) + (st_controllogic s)) + (create_reg_state_incr s). + +Definition transf_module (f: function) : mon module := + do fin <- create_reg; + do rtrn <- create_reg; + do _ <- traverselist (transf_instr fin rtrn) (Maps.PTree.elements f.(RTL.fn_code)); + do start <- create_reg; + do rst <- create_reg; + do clk <- create_reg; + do current_state <- get; + ret (mkmodule + f.(RTL.fn_params) + current_state.(st_datapath) + current_state.(st_controllogic) + f.(fn_entrypoint) + current_state.(st_st) + fin + rtrn). + +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)) + (st_datapath (init_state st)) + (st_controllogic (init_state st)). + +Definition transl_module (f : function) : Errors.res module := + run_mon (max_state f) (transf_module f). + +Fixpoint main_function (main : ident) (flist : list (ident * AST.globdef fundef unit)) +{struct flist} : option function := + match flist with + | (i, AST.Gfun (AST.Internal f)) :: xs => + if Pos.eqb i main + then Some f + else main_function main xs + | _ :: xs => main_function main xs + | nil => None + end. + +Definition transf_program (d : program) : Errors.res module := + match main_function d.(AST.prog_main) d.(AST.prog_defs) with + | Some f => run_mon (max_state f) (transf_module f) + | _ => Errors.Error (Errors.msg "HTL: could not find main function") + end. diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index 9caa2d1..cac2c02 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -43,15 +43,14 @@ Module AssocMapExt. forall am am' r v, merge (add r v am) am' = add r v (merge am am'). Proof. - intros. - unfold merge. - Search fold_right. - apply SetoidList.fold_right_add2. + induction am; intros. + unfold merge. simpl. unfold fold_right. simpl. Search MapsTo. + Abort. Lemma merge_base : forall am, merge (empty elt) am = am. - Proof. intros. unfold merge. apply fold_1. Qed. + Proof. auto. Qed. End Operations. -- cgit From 0c5ca8061a887f397991e62b580d5fc2f6484336 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 24 May 2020 13:37:26 +0100 Subject: Add .gitignore for *.vok *.vos --- .gitignore | 2 ++ 1 file changed, 2 insertions(+) diff --git a/.gitignore b/.gitignore index 7c17072..b2cf595 100644 --- a/.gitignore +++ b/.gitignore @@ -17,6 +17,8 @@ *.v.d *.vio *.vo +*.vok +*.vos .coq-native/ .csdp.cache .lia.cache -- cgit From 9f631b114e1cdcbaf3610eefd6f9fd7baeeda0c1 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 25 May 2020 22:17:11 +0100 Subject: Continuing work on proving specification --- src/common/Monad.v | 6 ++ src/translation/HTLgen.v | 13 ++- src/translation/HTLgenspec.v | 227 ++++++++++++++++++++++++++++++++++++++++--- 3 files changed, 224 insertions(+), 22 deletions(-) diff --git a/src/common/Monad.v b/src/common/Monad.v index 773918a..8517186 100644 --- a/src/common/Monad.v +++ b/src/common/Monad.v @@ -39,4 +39,10 @@ Module MonadExtra(M : Monad). ret (r::rs) end. + Fixpoint collectlist {A : Type} (f : A -> mon unit) (l : list A) {struct l} : mon unit := + match l with + | nil => ret tt + | x::xs => do _ <- f x; collectlist f xs + end. + End MonadExtra. diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 7bb5030..714cf98 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -63,8 +63,7 @@ Module HTLState <: State. Definition st_prop := st_incr. Hint Unfold st_prop : htlh. - Lemma st_refl : forall s, st_prop s s. - Proof. auto with htlh. Qed. + 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. @@ -85,14 +84,14 @@ Module HTLState <: State. Qed. End HTLState. -Import HTLState. +Export HTLState. Module HTLMonad := Statemonad(HTLState). -Import HTLMonad. +Export HTLMonad. Module HTLMonadExtra := Monad.MonadExtra(HTLMonad). Import HTLMonadExtra. -Import MonadNotation. +Export MonadNotation. Definition state_goto (st : reg) (n : node) : stmnt := Vnonblock (Vvar st) (Vlit (posToValue 32 n)). @@ -316,7 +315,7 @@ Definition create_reg: mon reg := Definition transf_module (f: function) : mon module := do fin <- create_reg; do rtrn <- create_reg; - do _ <- traverselist (transf_instr fin rtrn) (Maps.PTree.elements f.(RTL.fn_code)); + do _ <- collectlist (transf_instr fin rtrn) (Maps.PTree.elements f.(RTL.fn_code)); do start <- create_reg; do rst <- create_reg; do clk <- create_reg; @@ -354,6 +353,6 @@ Fixpoint main_function (main : ident) (flist : list (ident * AST.globdef fundef Definition transf_program (d : program) : Errors.res module := match main_function d.(AST.prog_main) d.(AST.prog_defs) with - | Some f => run_mon (max_state f) (transf_module f) + | Some f => transl_module f | _ => Errors.Error (Errors.msg "HTL: could not find main function") end. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index a7d65fc..8ea4384 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -17,8 +17,92 @@ *) From Coq Require Import FSets.FMapPositive. -From compcert Require RTL Op Maps. -From coqup Require Import Coquplib Verilog Veriloggen Value HTL. +From compcert Require RTL Op Maps Errors. +From coqup Require Import Coquplib Verilog Value HTL HTLgen AssocMap. + +Remark bind_inversion: + forall (A B: Type) (f: mon A) (g: A -> mon B) + (y: B) (s1 s3: st) (i: st_incr s1 s3), + bind f g s1 = OK y s3 i -> + exists x, exists s2, exists i1, exists i2, + f s1 = OK x s2 i1 /\ g x s2 = OK y s3 i2. +Proof. + intros until i. unfold bind. destruct (f s1); intros. + discriminate. + exists a; exists s'; exists s. + destruct (g a s'); inv H. + exists s0; auto. +Qed. + +Remark bind2_inversion: + forall (A B C: Type) (f: mon (A*B)) (g: A -> B -> mon C) + (z: C) (s1 s3: st) (i: st_incr s1 s3), + bind2 f g s1 = OK z s3 i -> + exists x, exists y, exists s2, exists i1, exists i2, + f s1 = OK (x, y) s2 i1 /\ g x y s2 = OK z s3 i2. +Proof. + unfold bind2; intros. + exploit bind_inversion; eauto. + intros [[x y] [s2 [i1 [i2 [P Q]]]]]. simpl in Q. + exists x; exists y; exists s2; exists i1; exists i2; auto. +Qed. + +Ltac monadInv1 H := + match type of H with + | (OK _ _ _ = OK _ _ _) => + inversion H; clear H; try subst + | (Error _ _ = OK _ _ _) => + discriminate + | (ret _ _ = OK _ _ _) => + inversion H; clear H; try subst + | (error _ _ = OK _ _ _) => + discriminate + | (bind ?F ?G ?S = OK ?X ?S' ?I) => + let x := fresh "x" in ( + let s := fresh "s" in ( + let i1 := fresh "INCR" in ( + let i2 := fresh "INCR" in ( + let EQ1 := fresh "EQ" in ( + let EQ2 := fresh "EQ" in ( + destruct (bind_inversion _ _ F G X S S' I H) as [x [s [i1 [i2 [EQ1 EQ2]]]]]; + clear H; + try (monadInv1 EQ2))))))) + | (bind2 ?F ?G ?S = OK ?X ?S' ?I) => + let x1 := fresh "x" in ( + let x2 := fresh "x" in ( + let s := fresh "s" in ( + let i1 := fresh "INCR" in ( + let i2 := fresh "INCR" in ( + let EQ1 := fresh "EQ" in ( + let EQ2 := fresh "EQ" in ( + destruct (bind2_inversion _ _ _ F G X S S' I H) as [x1 [x2 [s [i1 [i2 [EQ1 EQ2]]]]]]; + clear H; + try (monadInv1 EQ2)))))))) + end. + +Ltac monadInv H := + match type of H with + | (ret _ _ = OK _ _ _) => monadInv1 H + | (error _ _ = OK _ _ _) => monadInv1 H + | (bind ?F ?G ?S = OK ?X ?S' ?I) => monadInv1 H + | (bind2 ?F ?G ?S = OK ?X ?S' ?I) => monadInv1 H + | (?F _ _ _ _ _ _ _ _ = OK _ _ _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ _ _ = OK _ _ _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ _ = OK _ _ _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ _ = OK _ _ _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ _ = OK _ _ _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ _ = OK _ _ _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ _ = OK _ _ _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + | (?F _ = OK _ _ _) => + ((progress simpl in H) || unfold F in H); monadInv1 H + end. (** * Relational specification of the translation *) @@ -64,29 +148,142 @@ Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Pr translate_condition cond args s = OK c s' i -> tr_instr fin rtrn st (RTL.Icond cond args n1 n2) Vskip (state_cond st c n1 n2) | tr_instr_Ireturn_None : - tr_instr fin rtrn st (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vlit (ZToValue 1%nat 0%Z)))) Vskip + tr_instr fin rtrn st (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vlit (ZToValue 1%nat 0%Z)))) Vskip | tr_instr_Ireturn_Some : forall r, tr_instr fin rtrn st (RTL.Ireturn (Some r)) (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r))) Vskip. -Inductive tr_code (c : RTL.code) (pc : RTL.node) (stmnts trans : PositiveMap.t stmnt) +Inductive tr_code (pc : RTL.node) (i : RTL.instruction) (stmnts trans : PositiveMap.t stmnt) (fin rtrn st : reg) : Prop := tr_code_intro : - forall i s t, - Maps.PTree.get pc c = Some i -> + forall s t, stmnts!pc = Some s -> trans!pc = Some t -> tr_instr fin rtrn st i s t -> - tr_code c pc stmnts trans fin rtrn st. + tr_code pc i stmnts trans fin rtrn st. Inductive tr_module (f : RTL.function) : module -> Prop := tr_module_intro : - forall data control fin rtrn st, - (forall pc, tr_code f.(RTL.fn_code) pc data control fin rtrn st) -> - tr_module f (mkmodule - f.(RTL.fn_params) - data - control - f.(RTL.fn_entrypoint) - st fin rtrn). + forall data control fin rtrn st m, + (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i -> + tr_code pc i data control fin rtrn st) -> + m = (mkmodule f.(RTL.fn_params) + data + control + f.(RTL.fn_entrypoint) + st fin rtrn) -> + tr_module f m. + +Lemma create_reg_datapath_trans : + forall s s' x i, + create_reg s = OK x s' i -> + s.(st_datapath) = s'.(st_datapath). +Proof. intros. monadInv H. trivial. Qed. + +Lemma create_reg_controllogic_trans : + forall s s' x i, + create_reg s = OK x s' i -> + s.(st_controllogic) = s'.(st_controllogic). +Proof. intros. monadInv H. trivial. Qed. + +Lemma get_refl_x : + forall s s' x i, + get s = OK x s' i -> + s = x. +Proof. inversion 1. trivial. Qed. + +Lemma get_refl_s : + forall s s' x i, + get s = OK x s' i -> + s = s'. +Proof. inversion 1. trivial. Qed. + +Ltac inv_incr := + match goal with + | [ H: create_reg ?s = OK _ ?s' _ |- _ ] => + let H1 := fresh "H" in + assert (H1 := H); eapply create_reg_datapath_trans in H; + eapply create_reg_controllogic_trans in H1; inv_incr + | [ H: get ?s = OK _ _ _ |- _ ] => + let H1 := fresh "H" in + assert (H1 := H); apply get_refl_x in H; apply get_refl_s in H1; + subst; inv_incr + | [ H: st_prop _ _ |- _ ] => unfold st_prop in H; destruct H; inv_incr + | [ H: st_incr _ _ |- _ ] => destruct st_incr; inv_incr + | _ => idtac + end. + +Ltac rewrite_states := + match goal with + | [ H: ?x ?s = ?x ?s' |- _ ] => + let c1 := fresh "c" in + let c2 := fresh "c" in + remember (?x ?s) as c1; remember (?x ?s') as c2; try subst + end. + +Lemma iter_expand_instr_spec : + forall l fin rtrn s s' i x, + HTLMonadExtra.collectlist (transf_instr fin rtrn) l s = OK x s' i -> + list_norepet (List.map fst l) -> + (forall pc instr, In (pc, instr) l -> + tr_code pc instr s'.(st_datapath) s'.(st_controllogic) fin rtrn s'.(st_st)). +Proof. + induction l; simpl; intros. + - contradiction. + - destruct a as [pc1 instr1]; simpl in *. inv H0. monadInv H. inv_incr. + destruct (peq pc pc1). + + + subst. + destruct instr1. + econstructor. + unfold add_instr in EQ. + destruct (check_empty_node_datapath s1 pc1); try discriminate. + destruct (check_empty_node_controllogic s1 pc1); try discriminate. + inversion EQ. + destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + + unfold add_instr in EQ. + destruct (check_empty_node_datapath s1 pc1); try discriminate. + destruct (check_empty_node_controllogic s1 pc1); try discriminate. + inversion EQ. + destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + + + eapply IHl. apply EQ0. assumption. + destruct H1. inversion H1. subst. contradiction. + assumption. + +Qed. + +Theorem transl_module_correct : + forall f m, + transl_module f = Errors.OK m -> tr_module f m. +Proof. + intros until m. + unfold transl_module. + unfold run_mon. + destruct (transf_module f (max_state f)) eqn:?; try discriminate. + intros. inv H. + inversion s; subst. + + unfold transf_module in *. + monadInv Heqr. + econstructor; simpl; trivial. + intros. + inv_incr. + assert (st_controllogic s8 = st_controllogic s2). + { rewrite <- H5. rewrite <- H6. rewrite <- H7. trivial. } + rewrite <- H10. + assert (st_datapath s8 = st_datapath s2). + { rewrite <- EQ4. rewrite <- EQ3. rewrite <- EQ2. trivial. } + assert (st_st s5 = st_st s2). + { rewrite H10. rewrite <- H50. trivial. } + rewrite H80. rewrite H81. rewrite H82. + eapply iter_expand_instr_spec. + apply EQ0. + apply Maps.PTree.elements_keys_norepet. + apply Maps.PTree.elements_correct. + assumption. +Qed. -- cgit From 43b078ba35aec32840f459d67de3fe6408184cea Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 25 May 2020 22:32:33 +0100 Subject: Finished proving the first case --- src/translation/HTLgenspec.v | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 8ea4384..f01ab1c 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -235,7 +235,7 @@ Proof. destruct (peq pc pc1). + subst. - destruct instr1. + destruct instr1 eqn:?. econstructor. unfold add_instr in EQ. destruct (check_empty_node_datapath s1 pc1); try discriminate. @@ -251,10 +251,15 @@ Proof. destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. + inversion H1. inversion H7. rewrite H. apply tr_instr_Inop. + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + eapply IHl. apply EQ0. assumption. destruct H1. inversion H1. subst. contradiction. assumption. + + Qed. Theorem transl_module_correct : -- cgit From bdd3b6734690100a7c696cf57bfe52963ec2c6ef Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 26 May 2020 12:04:59 +0100 Subject: Finished second pass and fixed bug --- src/translation/HTLgen.v | 2 +- src/translation/HTLgenspec.v | 53 ++++++++++++++++++++++++++++++-------------- 2 files changed, 37 insertions(+), 18 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 714cf98..670dcd4 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -272,7 +272,7 @@ Definition transf_instr (fin rtrn: reg) (ni: node * instruction) : mon unit := | Inop n' => add_instr n n' Vskip | Iop op args dst n' => do instr <- translate_instr op args; - add_instr n n' (block dst instr) + add_instr n n' (nonblock dst instr) | Iload _ _ _ _ _ => error (Errors.msg "Loads are not implemented.") | Istore _ _ _ _ _ => error (Errors.msg "Stores are not implemented.") | Icall _ _ _ _ _ => error (Errors.msg "Calls are not implemented.") diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index f01ab1c..c8f1751 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -222,6 +222,12 @@ Ltac rewrite_states := remember (?x ?s) as c1; remember (?x ?s') as c2; try subst end. +Lemma translate_instr_tr_op : + forall op args s s' e i, + translate_instr op args s = OK e s' i -> + tr_op op args e. +Admitted. + Lemma iter_expand_instr_spec : forall l fin rtrn s s' i x, HTLMonadExtra.collectlist (transf_instr fin rtrn) l s = OK x s' i -> @@ -236,30 +242,43 @@ Proof. + subst. destruct instr1 eqn:?. - econstructor. - unfold add_instr in EQ. - destruct (check_empty_node_datapath s1 pc1); try discriminate. - destruct (check_empty_node_controllogic s1 pc1); try discriminate. - inversion EQ. - destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - unfold add_instr in EQ. - destruct (check_empty_node_datapath s1 pc1); try discriminate. - destruct (check_empty_node_controllogic s1 pc1); try discriminate. - inversion EQ. - destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. + * unfold add_instr in EQ. + destruct (check_empty_node_datapath s1 pc1); try discriminate. + destruct (check_empty_node_controllogic s1 pc1); try discriminate. + inversion EQ. + econstructor. + destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + + destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + + inversion H1. inversion H7. rewrite H. apply tr_instr_Inop. + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + * monadInv EQ. + inv_incr. + unfold add_instr in EQ2. + destruct (check_empty_node_datapath s0 pc1); try discriminate. + destruct (check_empty_node_controllogic s0 pc1); try discriminate. + inversion EQ2. + econstructor. + destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. - inversion H1. inversion H7. rewrite H. apply tr_instr_Inop. - eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + + inversion H1. inversion H7. unfold nonblock. rewrite <- e2. rewrite H. constructor. + eapply translate_instr_tr_op. apply EQ1. + + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + eapply IHl. apply EQ0. assumption. destruct H1. inversion H1. subst. contradiction. assumption. - - Qed. Theorem transl_module_correct : -- cgit From b3c58bddaa0aa06f2c0905eeb3eb5318cb6361aa Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 26 May 2020 13:02:15 +0100 Subject: Finished proof of spec completely --- src/translation/HTLgen.v | 35 ++++++++++++++++++++++-- src/translation/HTLgenspec.v | 64 +++++++++++++++++++++++++++++++++++++++++--- 2 files changed, 94 insertions(+), 5 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 670dcd4..26dc630 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -142,6 +142,37 @@ Definition add_instr (n : node) (n' : node) (st : stmnt) : mon unit := | _, _ => Error (Errors.msg "HTL.add_instr") end. +Lemma add_instr_skip_state_incr : + forall s n st, + (st_datapath s)!n = None -> + (st_controllogic s)!n = None -> + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + (AssocMap.add n st s.(st_datapath)) + (AssocMap.add 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_datapath s n, check_empty_node_controllogic s n with + | left STM, left TRANS => + OK tt (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + (AssocMap.add n st s.(st_datapath)) + (AssocMap.add n Vskip s.(st_controllogic))) + (add_instr_skip_state_incr s n st STM TRANS) + | _, _ => Error (Errors.msg "HTL.add_instr") + end. + Definition nonblock (dst : reg) (e : expr) := Vnonblock (Vvar dst) e. Definition block (dst : reg) (e : expr) := Vblock (Vvar dst) e. @@ -285,9 +316,9 @@ Definition transf_instr (fin rtrn: reg) (ni: node * instruction) : mon unit := | Ireturn r => match r with | Some r' => - add_instr n n (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r'))) + add_instr_skip n (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r'))) | None => - add_instr n n (block fin (Vlit (ZToValue 1%nat 1%Z))) + add_instr_skip n (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vlit (ZToValue 1%nat 0%Z)))) end end end. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index c8f1751..06ed68d 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -148,7 +148,8 @@ Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Pr translate_condition cond args s = OK c s' i -> tr_instr fin rtrn st (RTL.Icond cond args n1 n2) Vskip (state_cond st c n1 n2) | tr_instr_Ireturn_None : - tr_instr fin rtrn st (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vlit (ZToValue 1%nat 0%Z)))) Vskip + tr_instr fin rtrn st (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) + (block rtrn (Vlit (ZToValue 1%nat 0%Z)))) Vskip | tr_instr_Ireturn_Some : forall r, tr_instr fin rtrn st (RTL.Ireturn (Some r)) @@ -226,7 +227,15 @@ Lemma translate_instr_tr_op : forall op args s s' e i, translate_instr op args s = OK e s' i -> tr_op op args e. -Admitted. +Proof. + intros. + destruct op eqn:?; try (econstructor; apply H); try discriminate; simpl in *; + try (match goal with + [ H: match ?args with _ => _ end _ = _ _ _ |- _ ] => + repeat (destruct args; try discriminate) + end); + monadInv H; constructor. +Qed. Lemma iter_expand_instr_spec : forall l fin rtrn s s' i x, @@ -241,7 +250,7 @@ Proof. destruct (peq pc pc1). + subst. - destruct instr1 eqn:?. + destruct instr1 eqn:?; try discriminate. * unfold add_instr in EQ. destruct (check_empty_node_datapath s1 pc1); try discriminate. @@ -275,6 +284,55 @@ Proof. eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + * monadInv EQ. + inv_incr. + unfold add_branch_instr in EQ2. + destruct (check_empty_node_datapath s0 pc1); try discriminate. + destruct (check_empty_node_controllogic s0 pc1); try discriminate. + inversion EQ2. + + econstructor. + destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + + inversion H1. inversion H7. rewrite <- e2. rewrite H. econstructor. + apply EQ1. + + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + * destruct o3 eqn:?. + unfold add_instr_skip in EQ. + destruct (check_empty_node_datapath s1 pc1); try discriminate. + destruct (check_empty_node_controllogic s1 pc1); try discriminate. + inversion EQ. + + econstructor. + destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + + inversion H1. inversion H7. constructor. + + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + unfold add_instr_skip in EQ. + destruct (check_empty_node_datapath s1 pc1); try discriminate. + destruct (check_empty_node_controllogic s1 pc1); try discriminate. + inversion EQ. + + econstructor. + destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + + inversion H1. inversion H7. apply constructor. + + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + eapply IHl. apply EQ0. assumption. destruct H1. inversion H1. subst. contradiction. assumption. -- cgit From e4dac24479752b7d2645c9ba9dc23b73051f0f9e Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 26 May 2020 17:33:45 +0100 Subject: Working on automation --- src/translation/HTLgenspec.v | 110 +++++++++++++++++++------------------------ 1 file changed, 48 insertions(+), 62 deletions(-) diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 06ed68d..9c46b48 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -237,6 +237,37 @@ Proof. monadInv H; constructor. Qed. +Ltac unfold_match H := + match type of H with + | context[match ?g with _ => _ end] => destruct g; try discriminate + end. + +Ltac inv_add_instr' H := + match type of H with + | ?f _ _ _ = OK _ _ _ => unfold f in H + | ?f _ _ _ _ = OK _ _ _ => unfold f in H + | ?f _ _ _ _ _ = OK _ _ _ => unfold f in H + end; repeat unfold_match H; inversion H. + +Ltac inv_add_instr := + match goal with + | H: context[add_instr_skip _ _ _] |- _ => + inv_add_instr' H + | H: context[add_instr_skip _ _] |- _ => + monadInv H; inv_incr; inv_add_instr + | H: context[add_instr _ _ _ _] |- _ => + inv_add_instr' H + | H: context[add_instr _ _ _] |- _ => + monadInv H; inv_incr; inv_add_instr + | H: context[add_branch_instr _ _ _ _ _] |- _ => + inv_add_instr' H + | H: context[add_branch_instr _ _ _ _] |- _ => + monadInv H; inv_incr; inv_add_instr + end. + +Ltac destruct_optional := + match goal with H: option ?r |- _ => destruct H end. + Lemma iter_expand_instr_spec : forall l fin rtrn s s' i x, HTLMonadExtra.collectlist (transf_instr fin rtrn) l s = OK x s' i -> @@ -250,87 +281,42 @@ Proof. destruct (peq pc pc1). + subst. - destruct instr1 eqn:?; try discriminate. - - * unfold add_instr in EQ. - destruct (check_empty_node_datapath s1 pc1); try discriminate. - destruct (check_empty_node_controllogic s1 pc1); try discriminate. - inversion EQ. - econstructor. - destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. + destruct instr1 eqn:?; try discriminate; try destruct_optional; inv_add_instr; econstructor. - destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - - inversion H1. inversion H7. rewrite H. apply tr_instr_Inop. + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * inversion H1. inversion H7. rewrite H. apply tr_instr_Inop. eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - * monadInv EQ. - inv_incr. - unfold add_instr in EQ2. - destruct (check_empty_node_datapath s0 pc1); try discriminate. - destruct (check_empty_node_controllogic s0 pc1); try discriminate. - inversion EQ2. - econstructor. - destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - - destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - - inversion H1. inversion H7. unfold nonblock. rewrite <- e2. rewrite H. constructor. + * inversion H1. inversion H7. unfold nonblock. rewrite <- e2. rewrite H. constructor. eapply translate_instr_tr_op. apply EQ1. - eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - * monadInv EQ. - inv_incr. - unfold add_branch_instr in EQ2. - destruct (check_empty_node_datapath s0 pc1); try discriminate. - destruct (check_empty_node_controllogic s0 pc1); try discriminate. - inversion EQ2. - - econstructor. - destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - - inversion H1. inversion H7. rewrite <- e2. rewrite H. econstructor. - apply EQ1. - + * inversion H1. inversion H7. rewrite <- e2. rewrite H. econstructor. apply EQ1. eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - * destruct o3 eqn:?. - unfold add_instr_skip in EQ. - destruct (check_empty_node_datapath s1 pc1); try discriminate. - destruct (check_empty_node_controllogic s1 pc1); try discriminate. - inversion EQ. - - econstructor. - destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - - inversion H1. inversion H7. constructor. - + * inversion H1. inversion H7. constructor. eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - unfold add_instr_skip in EQ. - destruct (check_empty_node_datapath s1 pc1); try discriminate. - destruct (check_empty_node_controllogic s1 pc1); try discriminate. - inversion EQ. - - econstructor. - destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; [ discriminate | apply H7 ]. - - inversion H1. inversion H7. apply constructor. - + * inversion H1. inversion H7. constructor. eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + eapply IHl. apply EQ0. assumption. -- cgit From cb00777e134464a01a9e97efb448cee7473df9d5 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 27 May 2020 15:05:20 +0100 Subject: Add top level definition --- src/translation/HTLgenproof.v | 166 ++++++++++++++++++++++-------------------- src/translation/HTLgenspec.v | 125 ++++++++++++++++--------------- 2 files changed, 153 insertions(+), 138 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index cb621a8..b07c654 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -16,34 +16,39 @@ * along with this program. If not, see . *) -From coqup Require Import HTLgenspec Value AssocMap. +From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap. From coqup Require HTL Verilog. From compcert Require RTL Registers Globalenvs AST. Import AssocMapNotation. +Hint Resolve Smallstep.forward_simulation_plus : htlproof. +Hint Resolve AssocMap.gss : htlproof. +Hint Resolve AssocMap.gso : htlproof. + Inductive match_assocmaps : RTL.regset -> assocmap -> Prop := | match_assocmap : forall rs am, (forall r, val_value_lessdef (Registers.Regmap.get r rs) am#r) -> match_assocmaps rs am. +Hint Constructors match_assocmaps : htlproof. Definition state_st_wf (m : HTL.module) (s : HTL.state) := forall st assoc, s = HTL.State m st assoc -> assoc!(m.(HTL.mod_st)) = Some (posToValue 32 st). +Hint Unfold state_st_wf : htlproof. Inductive match_states : RTL.state -> HTL.state -> Prop := | match_state : forall (rs : RTL.regset) assoc sf f sp rs mem m st (MASSOC : match_assocmaps rs assoc) (TF : tr_module f m) - (TC : tr_code f.(RTL.fn_code) st m.(HTL.mod_datapath) m.(HTL.mod_controllogic) - m.(HTL.mod_finish) m.(HTL.mod_return) m.(HTL.mod_st)) (WF : state_st_wf m (HTL.State m st assoc)), match_states (RTL.State sf f sp st rs mem) (HTL.State m st assoc) | match_returnstate : forall v v' stack m, val_value_lessdef v v' -> match_states (RTL.Returnstate stack v m) (HTL.Returnstate v'). +Hint Constructors match_states : htlproof. Inductive match_program : RTL.program -> HTL.module -> Prop := match_program_intro : @@ -53,6 +58,7 @@ Inductive match_program : RTL.program -> HTL.module -> Prop := Globalenvs.Genv.find_funct_ptr ge b = Some (AST.Internal f) -> tr_module f m -> match_program p m. +Hint Constructors match_program : htlproof. Lemma regs_lessdef_regs : forall rs1 rs2 n v, @@ -77,11 +83,6 @@ Lemma valToValue_lessdef : val_value_lessdef v (valToValue v). Admitted. -Lemma assocmap_merge_add : - forall k v assoc, - AssocMap.add k v assoc = merge_assocmap (AssocMap.add k v empty_assocmap) assoc. -Admitted. - (* Need to eventually move posToValue 32 to posToValueAuto, as that has this proof. *) Lemma assumption_32bit : forall v, @@ -116,30 +117,43 @@ Lemma regset_assocmap_wf2 : Admitted. Lemma st_greater_than_res : - forall m res, - HTL.mod_st m <> res. + forall m res : positive, + m <> res. Admitted. Lemma finish_not_return : - forall m, - HTL.mod_finish m <> HTL.mod_return m. + forall r f : positive, + r <> f. Admitted. Lemma finish_not_res : - forall m r, - HTL.mod_finish m <> r. + forall f r : positive, + f <> r. Admitted. -Ltac subst_eq_hyp := - match goal with - H1: ?x = Some ?i, H2: ?x = Some ?j |- _ => - let H := fresh "H" in - rewrite H1 in H2; injection H2; intro H; clear H2; subst - end. - Ltac unfold_merge := try (repeat (rewrite merge_inj)); unfold merge_assocmap; rewrite AssocMapExt.merge_base. +Ltac inv_state := + match goal with + MSTATE : match_states _ _ |- _ => + inversion MSTATE; + match goal with + TF : tr_module _ _ |- _ => + inversion TF; + match goal with + TC : forall _ _, + Maps.PTree.get _ _ = Some _ -> tr_code _ _ _ _ _ _ _, + H : Maps.PTree.get _ _ = Some _ |- _ => + apply TC in H; inversion H; + match goal with + TI : tr_instr _ _ _ _ _ _ |- _ => + inversion TI + end + end + end + end; subst. + Section CORRECTNESS. Variable prog : RTL.program. @@ -160,7 +174,7 @@ Section CORRECTNESS. Lemma eval_cond_correct : forall cond (args : list Registers.reg) s1 c s' i rs args m b f assoc, - Veriloggen.translate_condition cond args s1 = Veriloggen.OK c s' i -> + translate_condition cond args s1 = OK c s' i -> Op.eval_condition cond (List.map (fun r : BinNums.positive => Registers.Regmap.get r rs) args) m = Some b -> Verilog.expr_runp f assoc c (boolToValue 32 b). Admitted. @@ -195,56 +209,37 @@ Section CORRECTNESS. match_states S1 R1 -> exists R2, Smallstep.plus HTL.step tge R1 t R2 /\ match_states S2 R2. Proof. - induction 1; intros R1 MSTATE; - try (inversion MSTATE; inversion TC; inversion H12; subst_eq_hyp; discriminate). + induction 1; intros R1 MSTATE; try inv_state. - (* Inop *) - inversion MSTATE. subst. econstructor. split. apply Smallstep.plus_one. - inversion TC. - eapply HTL.step_module; subst_eq_hyp; inversion H3; subst. - apply H2. - apply H1. + eapply HTL.step_module; eauto. (* processing of state *) econstructor. simpl. trivial. econstructor. trivial. - inversion H3. subst. econstructor. - (* prove merge *) - apply assocmap_merge_add. - (* prove stval *) + unfold_merge; simpl. apply AssocMap.gss. - apply assumption_32bit. (* prove match_state *) - constructor. - apply rs. - apply regs_lessdef_regs. - assumption. - assumption. - inversion TF. simpl. apply H0. - unfold state_st_wf. intros. inversion H0. subst. - apply AssocMap.gss. + unfold_merge. rewrite assumption_32bit. + constructor; auto. + apply regs_lessdef_regs. assumption. + unfold state_st_wf. inversion 1. subst. apply AssocMap.gss. - (* Iop *) - inversion MSTATE. subst. econstructor. split. apply Smallstep.plus_one. - inversion TC. - eapply HTL.step_module; subst_eq_hyp; inversion H4; subst. - apply H3. - apply H2. - econstructor. - simpl. trivial. - constructor. trivial. - econstructor. - simpl. trivial. - eapply eval_correct. apply H0. apply H10. trivial. trivial. - unfold_merge. + eapply HTL.step_module; eauto. + econstructor; simpl; trivial. + constructor; trivial. + econstructor; simpl; eauto. + eapply eval_correct; eauto. + unfold_merge. simpl. rewrite AssocMap.gso. apply AssocMap.gss. apply st_greater_than_res. @@ -260,30 +255,24 @@ Section CORRECTNESS. assumption. apply valToValue_lessdef. assumption. - inversion TF. simpl. apply H2. unfold state_st_wf. intros. inversion H2. subst. unfold_merge. rewrite AssocMap.gso. apply AssocMap.gss. apply st_greater_than_res. - - inversion MSTATE; inversion TC; - inversion H12; subst_eq_hyp; inversion H13. - - inversion MSTATE. inversion TC. subst_eq_hyp. inversion H12. subst. - econstructor. split. apply Smallstep.plus_one. - eapply HTL.step_module. subst. apply H11. apply H10. + - econstructor. split. apply Smallstep.plus_one. + eapply HTL.step_module; eauto. eapply Verilog.stmnt_runp_Vnonblock with (rhsval := if b then posToValue 32 ifso else posToValue 32 ifnot). simpl. trivial. destruct b. eapply Verilog.erun_Vternary_true. - eapply eval_cond_correct. apply args. - apply H7. apply H0. + eapply eval_cond_correct; eauto. constructor. apply assumption_32bit_bool. eapply Verilog.erun_Vternary_false. - eapply eval_cond_correct. apply args. - apply H7. apply H0. + eapply eval_cond_correct; eauto. constructor. apply assumption_32bit_bool. trivial. @@ -300,8 +289,6 @@ Section CORRECTNESS. apply regs_lessdef_regs. assumption. assumption. - inversion TF. simpl. apply H1. - unfold state_st_wf. intros. inversion H1. subst. unfold_merge. apply AssocMap.gss. @@ -311,21 +298,15 @@ Section CORRECTNESS. apply regs_lessdef_regs. assumption. assumption. - inversion TF. simpl. apply H1. - unfold state_st_wf. intros. inversion H1. subst. unfold_merge. apply AssocMap.gss. - (* Return *) - inversion MSTATE. inversion TC. subst_eq_hyp. - inversion H11. subst. econstructor. split. eapply Smallstep.plus_two. - eapply HTL.step_module. - apply H10. - apply H9. + eapply HTL.step_module; eauto. constructor. econstructor; simpl; trivial. econstructor; simpl; trivial. @@ -340,6 +321,7 @@ Section CORRECTNESS. apply st_greater_than_res. apply st_greater_than_res. trivial. apply HTL.step_finish. + unfold_merge; simpl. rewrite AssocMap.gso. apply AssocMap.gss. apply finish_not_return. @@ -347,13 +329,11 @@ Section CORRECTNESS. rewrite Events.E0_left. trivial. constructor. constructor. - destruct (assoc!r) eqn:?. + - destruct (assoc!r) eqn:?. inversion H11. subst. econstructor. split. eapply Smallstep.plus_two. - eapply HTL.step_module. - apply H10. - apply H9. + eapply HTL.step_module; eauto. constructor. econstructor; simpl; trivial. econstructor; simpl; trivial. @@ -370,12 +350,42 @@ Section CORRECTNESS. apply st_greater_than_res. apply st_greater_than_res. trivial. apply HTL.step_finish. + unfold_merge. rewrite AssocMap.gso. apply AssocMap.gss. apply finish_not_return. apply AssocMap.gss. rewrite Events.E0_left. trivial. - constructor. simpl. Search Registers.Regmap.get. + constructor. simpl. + Admitted. + Hint Resolve transl_step_correct : htlproof. + + Lemma senv_preserved : + forall id : AST.ident, + Globalenvs.Senv.public_symbol (Smallstep.symbolenv (HTL.semantics tprog)) id = + Globalenvs.Senv.public_symbol (Smallstep.symbolenv (RTL.semantics prog)) id. + Proof. Admitted. + Hint Resolve senv_preserved : htlproof. + + Lemma transl_initial_states : + forall s1 : Smallstep.state (RTL.semantics prog), + Smallstep.initial_state (RTL.semantics prog) s1 -> + exists s2 : Smallstep.state (HTL.semantics tprog), + Smallstep.initial_state (HTL.semantics tprog) s2 /\ match_states s1 s2. + Proof. Admitted. + Hint Resolve transl_initial_states : htlproof. + + Lemma transl_final_states : + forall (s1 : Smallstep.state (RTL.semantics prog)) (s2 : Smallstep.state (HTL.semantics tprog)) + (r : Integers.Int.int), + match_states s1 s2 -> + Smallstep.final_state (RTL.semantics prog) s1 r -> Smallstep.final_state (HTL.semantics tprog) s2 r. + Proof. Admitted. + Hint Resolve transl_final_states : htlproof. + + Theorem transf_program_correct: + Smallstep.forward_simulation (RTL.semantics prog) (HTL.semantics tprog). + Proof. eauto with htlproof. Qed. End CORRECTNESS. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index 9c46b48..b9d483c 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -20,6 +20,9 @@ From Coq Require Import FSets.FMapPositive. From compcert Require RTL Op Maps Errors. From coqup Require Import Coquplib Verilog Value HTL HTLgen AssocMap. +Hint Resolve Maps.PTree.elements_keys_norepet : htlspec. +Hint Resolve Maps.PTree.elements_correct : htlspec. + Remark bind_inversion: forall (A B: Type) (f: mon A) (g: A -> mon B) (y: B) (s1 s3: st) (i: st_incr s1 s3), @@ -134,6 +137,7 @@ Inductive tr_op : Op.operation -> list reg -> expr -> Prop := | tr_op_Oshrimm : forall n r, tr_op (Op.Oshrimm n) (r::nil) (boplit Vshr r n) | tr_op_Ocmp : forall c l e s s' i, translate_condition c l s = OK e s' i -> tr_op (Op.Ocmp c) l e | tr_op_Olea : forall a l e s s' i, translate_eff_addressing a l s = OK e s' i -> tr_op (Op.Olea a) l e. +Hint Constructors tr_op : htlspec. Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Prop := | tr_instr_Inop : @@ -154,6 +158,7 @@ Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Pr forall r, tr_instr fin rtrn st (RTL.Ireturn (Some r)) (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r))) Vskip. +Hint Constructors tr_instr : htlspec. Inductive tr_code (pc : RTL.node) (i : RTL.instruction) (stmnts trans : PositiveMap.t stmnt) (fin rtrn st : reg) : Prop := @@ -163,6 +168,7 @@ Inductive tr_code (pc : RTL.node) (i : RTL.instruction) (stmnts trans : Positive trans!pc = Some t -> tr_instr fin rtrn st i s t -> tr_code pc i stmnts trans fin rtrn st. +Hint Constructors tr_code : htlspec. Inductive tr_module (f : RTL.function) : module -> Prop := tr_module_intro : @@ -175,44 +181,48 @@ Inductive tr_module (f : RTL.function) : module -> Prop := f.(RTL.fn_entrypoint) st fin rtrn) -> tr_module f m. +Hint Constructors tr_module : htlspec. Lemma create_reg_datapath_trans : forall s s' x i, create_reg s = OK x s' i -> s.(st_datapath) = s'.(st_datapath). Proof. intros. monadInv H. trivial. Qed. +Hint Resolve create_reg_datapath_trans : htlspec. Lemma create_reg_controllogic_trans : forall s s' x i, create_reg s = OK x s' i -> s.(st_controllogic) = s'.(st_controllogic). Proof. intros. monadInv H. trivial. Qed. +Hint Resolve create_reg_controllogic_trans : htlspec. Lemma get_refl_x : forall s s' x i, get s = OK x s' i -> s = x. Proof. inversion 1. trivial. Qed. +Hint Resolve get_refl_x : htlspec. Lemma get_refl_s : forall s s' x i, get s = OK x s' i -> s = s'. Proof. inversion 1. trivial. Qed. +Hint Resolve get_refl_s : htlspec. Ltac inv_incr := - match goal with + repeat match goal with | [ H: create_reg ?s = OK _ ?s' _ |- _ ] => let H1 := fresh "H" in assert (H1 := H); eapply create_reg_datapath_trans in H; - eapply create_reg_controllogic_trans in H1; inv_incr + eapply create_reg_controllogic_trans in H1 | [ H: get ?s = OK _ _ _ |- _ ] => let H1 := fresh "H" in assert (H1 := H); apply get_refl_x in H; apply get_refl_s in H1; - subst; inv_incr - | [ H: st_prop _ _ |- _ ] => unfold st_prop in H; destruct H; inv_incr - | [ H: st_incr _ _ |- _ ] => destruct st_incr; inv_incr - | _ => idtac + subst + | [ H: st_prop _ _ |- _ ] => unfold st_prop in H; destruct H + | [ H: st_incr _ _ |- _ ] => destruct st_incr end. Ltac rewrite_states := @@ -229,13 +239,14 @@ Lemma translate_instr_tr_op : tr_op op args e. Proof. intros. - destruct op eqn:?; try (econstructor; apply H); try discriminate; simpl in *; + destruct op eqn:?; eauto with htlspec; try discriminate; simpl in *; try (match goal with [ H: match ?args with _ => _ end _ = _ _ _ |- _ ] => repeat (destruct args; try discriminate) end); monadInv H; constructor. Qed. +Hint Resolve translate_instr_tr_op : htlspec. Ltac unfold_match H := match type of H with @@ -275,55 +286,53 @@ Lemma iter_expand_instr_spec : (forall pc instr, In (pc, instr) l -> tr_code pc instr s'.(st_datapath) s'.(st_controllogic) fin rtrn s'.(st_st)). Proof. - induction l; simpl; intros. - - contradiction. - - destruct a as [pc1 instr1]; simpl in *. inv H0. monadInv H. inv_incr. - destruct (peq pc pc1). - - + subst. - destruct instr1 eqn:?; try discriminate; try destruct_optional; inv_add_instr; econstructor. - - * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * inversion H1. inversion H7. rewrite H. apply tr_instr_Inop. - eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - - * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * inversion H1. inversion H7. unfold nonblock. rewrite <- e2. rewrite H. constructor. - eapply translate_instr_tr_op. apply EQ1. - eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - - * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * inversion H1. inversion H7. rewrite <- e2. rewrite H. econstructor. apply EQ1. - eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - - * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * inversion H1. inversion H7. constructor. - eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - - * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; - [ discriminate | apply H7 ]. - * inversion H1. inversion H7. constructor. - eapply in_map with (f := fst) in H7. simpl in H7. contradiction. - - + eapply IHl. apply EQ0. assumption. - destruct H1. inversion H1. subst. contradiction. - assumption. - + induction l; simpl; intros; try contradiction. + destruct a as [pc1 instr1]; simpl in *. inv H0. monadInv H. inv_incr. + destruct (peq pc pc1). + + subst. + destruct instr1 eqn:?; try discriminate; try destruct_optional; inv_add_instr; econstructor. + + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * inversion H1. inversion H7. rewrite H. apply tr_instr_Inop. + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * inversion H1. inversion H7. unfold nonblock. rewrite <- e2. rewrite H. constructor. + eapply translate_instr_tr_op. apply EQ1. + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * inversion H1. inversion H7. rewrite <- e2. rewrite H. econstructor. apply EQ1. + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * inversion H1. inversion H7. constructor. + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + * destruct o with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * destruct o0 with pc1; destruct H9; simpl in *; rewrite AssocMap.gss in H7; + [ discriminate | apply H7 ]. + * inversion H1. inversion H7. constructor. + eapply in_map with (f := fst) in H7. simpl in H7. contradiction. + + + eapply IHl. apply EQ0. assumption. + destruct H1. inversion H1. subst. contradiction. + assumption. Qed. +Hint Resolve iter_expand_instr_spec : htlspec. Theorem transl_module_correct : forall f m, @@ -349,9 +358,5 @@ Proof. assert (st_st s5 = st_st s2). { rewrite H10. rewrite <- H50. trivial. } rewrite H80. rewrite H81. rewrite H82. - eapply iter_expand_instr_spec. - apply EQ0. - apply Maps.PTree.elements_keys_norepet. - apply Maps.PTree.elements_correct. - assumption. + eauto with htlspec. Qed. -- cgit From 213267be33d714dabd19ca09210b5dc1ad4f6254 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 27 May 2020 23:38:40 +0100 Subject: Add more proofs and remove Admitted --- src/translation/HTLgenproof.v | 92 ++++++++++++++++++++----------------------- src/verilog/AssocMap.v | 45 +++++++++++++++++---- src/verilog/Value.v | 22 +++++++---- 3 files changed, 94 insertions(+), 65 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index b07c654..499b6b5 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -26,10 +26,11 @@ Hint Resolve Smallstep.forward_simulation_plus : htlproof. Hint Resolve AssocMap.gss : htlproof. Hint Resolve AssocMap.gso : htlproof. -Inductive match_assocmaps : RTL.regset -> assocmap -> Prop := -| match_assocmap : forall rs am, - (forall r, val_value_lessdef (Registers.Regmap.get r rs) am#r) -> - match_assocmaps rs am. +Inductive match_assocmaps : RTL.function -> RTL.regset -> assocmap -> Prop := + match_assocmap : forall f rs am, + (forall r, Ple r (RTL.max_reg_function f) -> + val_value_lessdef (Registers.Regmap.get r rs) am#r) -> + match_assocmaps f rs am. Hint Constructors match_assocmaps : htlproof. Definition state_st_wf (m : HTL.module) (s : HTL.state) := @@ -40,7 +41,7 @@ Hint Unfold state_st_wf : htlproof. Inductive match_states : RTL.state -> HTL.state -> Prop := | match_state : forall (rs : RTL.regset) assoc sf f sp rs mem m st - (MASSOC : match_assocmaps rs assoc) + (MASSOC : match_assocmaps f rs assoc) (TF : tr_module f m) (WF : state_st_wf m (HTL.State m st assoc)), match_states (RTL.State sf f sp st rs mem) @@ -61,27 +62,45 @@ Inductive match_program : RTL.program -> HTL.module -> Prop := Hint Constructors match_program : htlproof. Lemma regs_lessdef_regs : - forall rs1 rs2 n v, - match_assocmaps rs1 rs2 -> - match_assocmaps rs1 (AssocMap.add n v rs2). -Admitted. + forall f rs1 rs2 n v, + Plt (RTL.max_reg_function f) n -> + match_assocmaps f rs1 rs2 -> + match_assocmaps f rs1 (AssocMap.add n v rs2). +Proof. + inversion 2; subst. + intros. constructor. + intros. unfold find_assocmap. unfold AssocMapExt.find_default. + rewrite AssocMap.gso. eauto. + apply Pos.le_lt_trans with _ _ n in H2. + unfold not. intros. subst. eapply Pos.lt_irrefl. eassumption. assumption. +Qed. Lemma regs_add_match : - forall rs am r v v', - match_assocmaps rs am -> + forall f rs am r v v', val_value_lessdef v v' -> - match_assocmaps (Registers.Regmap.set r v rs) (AssocMap.add r v' am). -Admitted. - -Lemma merge_inj : - forall am am' r v, - merge_assocmap (AssocMap.add r v am) am' = AssocMap.add r v (merge_assocmap am am'). -Admitted. + match_assocmaps f rs am -> + match_assocmaps f (Registers.Regmap.set r v rs) (AssocMap.add r v' am). +Proof. + inversion 2; subst. + constructor. intros. + destruct (peq r0 r); subst. + rewrite Registers.Regmap.gss. + unfold find_assocmap. unfold AssocMapExt.find_default. + rewrite AssocMap.gss. assumption. + + rewrite Registers.Regmap.gso; try assumption. + unfold find_assocmap. unfold AssocMapExt.find_default. + rewrite AssocMap.gso; eauto. +Qed. Lemma valToValue_lessdef : - forall v, - val_value_lessdef v (valToValue v). -Admitted. + forall v v', + valToValue v = Some v' <-> + val_value_lessdef v v'. +Proof. + split; intros. + destruct v; try discriminate; constructor. + unfold valToValue in H. inversion H. unfold intToValue (* Need to eventually move posToValue 32 to posToValueAuto, as that has this proof. *) Lemma assumption_32bit : @@ -94,28 +113,6 @@ Lemma assumption_32bit_bool : valueToBool (boolToValue 32 b) = b. Admitted. -Lemma regset_assocmap_get : - forall r rs am v v', - match_assocmaps rs am -> - v = Registers.Regmap.get r rs -> - v' = am#r -> - val_value_lessdef v v'. -Proof. inversion 1. intros. subst. apply H0. Qed. - -Lemma regset_assocmap_wf : - forall r rs am i, - match_assocmaps rs am -> - Values.Vint i <> Registers.Regmap.get r rs -> - am!r = None. -Admitted. - -Lemma regset_assocmap_wf2 : - forall r rs am i, - match_assocmaps rs am -> - Values.Vint i = Registers.Regmap.get r rs -> - am!r = Some (intToValue i). -Admitted. - Lemma st_greater_than_res : forall m res : positive, m <> res. @@ -131,9 +128,6 @@ Lemma finish_not_res : f <> r. Admitted. -Ltac unfold_merge := - try (repeat (rewrite merge_inj)); unfold merge_assocmap; rewrite AssocMapExt.merge_base. - Ltac inv_state := match goal with MSTATE : match_states _ _ |- _ => @@ -222,12 +216,12 @@ Section CORRECTNESS. econstructor. (* prove stval *) - unfold_merge; simpl. - apply AssocMap.gss. + unfold merge_assocmap. apply AssocMapExt.merge_add. simpl. apply AssocMap.gss. (* prove match_state *) - unfold_merge. rewrite assumption_32bit. + rewrite assumption_32bit. constructor; auto. + unfold merge_assocmap. rewrite AssocMapExt.merge_add. simpl. apply AssocMap.gss. apply regs_lessdef_regs. assumption. unfold state_st_wf. inversion 1. subst. apply AssocMap.gss. - (* Iop *) diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index cac2c02..7db800b 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -26,6 +26,11 @@ Module AssocMap := PositiveMap. Module AssocMapExt. Import AssocMap. + Hint Resolve elements_3w : assocmap. + Hint Resolve elements_correct : assocmap. + Hint Resolve gss : assocmap. + Hint Resolve gso : assocmap. + Section Operations. Variable elt : Type. @@ -36,16 +41,40 @@ Module AssocMapExt. | Some v => v end. - Definition merge (m1 m2 : t elt) : t elt := - fold_right (fun el m => match el with (k, v) => add k v m end) m2 (elements m1). + Definition merge (am bm : t elt) : t elt := + fold_right (fun p a => add (fst p) (snd p) a) bm (elements am). - Lemma merge_add_assoc2 : - forall am am' r v, - merge (add r v am) am' = add r v (merge am am'). + Lemma add_assoc : + forall k v l bm, + List.In (k, v) l -> + SetoidList.NoDupA (@eq_key elt) l -> + @find elt k (fold_right (fun p a => add (fst p) (snd p) a) bm l) = Some v. Proof. - induction am; intros. - unfold merge. simpl. unfold fold_right. simpl. Search MapsTo. - Abort. + Hint Resolve SetoidList.InA_alt : assocmap. + Hint Extern 1 (exists _, _) => apply list_in_map_inv : assocmap. + + induction l; intros. + - contradiction. + - destruct a as [k' v']. + destruct (peq k k'). + + inversion H. inversion H1. inversion H0. subst. + simpl. auto with assocmap. + + subst. inversion H0. subst. apply in_map with (f := fst) in H1. simpl in *. + unfold not in H4. exfalso. apply H4. apply SetoidList.InA_alt. + auto with assocmap. + + + inversion H. inversion H1. inversion H0. subst. simpl. rewrite gso; try assumption. + apply IHl. contradiction. contradiction. + simpl. rewrite gso; try assumption. apply IHl. assumption. inversion H0. subst. assumption. + Qed. + Hint Resolve add_assoc : assocmap. + + Lemma merge_add : + forall k v am bm, + find k am = Some v -> + find k (merge am bm) = Some v. + Proof. unfold merge. auto with assocmap. Qed. Lemma merge_base : forall am, diff --git a/src/verilog/Value.v b/src/verilog/Value.v index 0ce0bc5..34cb0d2 100644 --- a/src/verilog/Value.v +++ b/src/verilog/Value.v @@ -85,10 +85,11 @@ Definition intToValue (i : Integers.int) : value := Definition valueToInt (i : value) : Integers.int := Int.repr (valueToZ i). -Definition valToValue (v : Values.val) := +Definition valToValue (v : Values.val) : option value := match v with - | Values.Vint i => intToValue i - | _ => ZToValue 32 0%Z + | Values.Vint i => Some (intToValue i) + | Values.Vundef => Some (ZToValue 32 0%Z) + | _ => None end. (** Convert a [value] to a [bool], so that choices can be made based on the @@ -295,17 +296,22 @@ Inductive val_value_lessdef: val -> value -> Prop := val_value_lessdef (Vint i) v' | lessdef_undef: forall v, val_value_lessdef Vundef v. -Search Z positive. - -Search "wordToZ". +Lemma valueToZ_ZToValue : + forall n z, + (- Z.of_nat (2 ^ n) <= z < Z.of_nat (2 ^ n))%Z -> + valueToZ (ZToValue (S n) z) = z. +Proof. + unfold valueToZ, ZToValue. simpl. + auto using wordToZ_ZToWord. +Qed. Lemma uvalueToZ_ZToValue : forall n z, (0 <= z < 2 ^ Z.of_nat n)%Z -> uvalueToZ (ZToValue n z) = z. Proof. - intros. unfold uvalueToZ, ZToValue. simpl. - apply uwordToZ_ZToWord. apply H. + unfold uvalueToZ, ZToValue. simpl. + auto using uwordToZ_ZToWord. Qed. Lemma valueToPos_posToValueAuto : -- cgit From 3ee0bfa7367fde5f9ec0dfd77f921f22f9e64563 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 28 May 2020 21:32:11 +0100 Subject: Finish Assocmap proofs --- src/verilog/AssocMap.v | 59 ++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 59 insertions(+) diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index 7db800b..39fe3d2 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -19,6 +19,10 @@ From coqup Require Import Coquplib Value. From Coq Require Import FSets.FMapPositive. +From compcert Require Import Maps. + +Search Maps.PTree.get not List.In. + Definition reg := positive. Module AssocMap := PositiveMap. @@ -70,16 +74,71 @@ Module AssocMapExt. Qed. Hint Resolve add_assoc : assocmap. + Lemma not_in_assoc : + forall k v l bm, + ~ List.In k (List.map (@fst key elt) l) -> + @find elt k bm = Some v -> + @find elt k (fold_right (fun p a => add (fst p) (snd p) a) bm l) = Some v. + Proof. + induction l; intros. + - assumption. + - destruct a as [k' v']. + destruct (peq k k'); subst; + simpl in *; apply Decidable.not_or in H; destruct H. contradiction. + rewrite AssocMap.gso; auto. + Qed. + Hint Resolve not_in_assoc : assocmap. + + Lemma elements_iff : + forall am k, + (exists v, find k am = Some v) <-> + List.In k (List.map (@fst _ elt) (elements am)). + Proof. + split; intros. + destruct H. apply elements_correct in H. apply in_map with (f := fst) in H. apply H. + apply list_in_map_inv in H. destruct H. destruct H. subst. + exists (snd x). apply elements_complete. assert (x = (fst x, snd x)) by apply surjective_pairing. + rewrite H in H0; assumption. + Qed. + Hint Resolve elements_iff : assocmap. + + Lemma elements_correct' : + forall am k, + ~ (exists v, find k am = Some v) <-> + ~ List.In k (List.map (@fst _ elt) (elements am)). + Proof. auto using not_iff_compat with assocmap. Qed. + Hint Resolve elements_correct' : assocmap. + + Lemma elements_correct_none : + forall am k, + find k am = None -> + ~ List.In k (List.map (@fst _ elt) (elements am)). + Proof. + intros. apply elements_correct'. unfold not. intros. + destruct H0. rewrite H in H0. discriminate. + Qed. + Hint Resolve elements_correct_none : assocmap. + Lemma merge_add : forall k v am bm, find k am = Some v -> find k (merge am bm) = Some v. Proof. unfold merge. auto with assocmap. Qed. + Hint Resolve merge_add : assocmap. + + Lemma merge_not_in : + forall k v am bm, + find k am = None -> + find k bm = Some v -> + find k (merge am bm) = Some v. + Proof. intros. apply not_in_assoc; auto with assocmap. Qed. + Hint Resolve merge_not_in : assocmap. Lemma merge_base : forall am, merge (empty elt) am = am. Proof. auto. Qed. + Hint Resolve merge_base : assocmap. End Operations. -- cgit From 3beb6e00213c5e7409d4574b663e5c6bf7490d66 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 29 May 2020 14:42:09 +0100 Subject: Change AssocMap to Maps.PTree --- src/translation/HTLgenproof.v | 8 +-- src/verilog/AssocMap.v | 113 ++++++++++++++++++++++++------------------ 2 files changed, 71 insertions(+), 50 deletions(-) diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index 499b6b5..b7b9d27 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -100,7 +100,8 @@ Lemma valToValue_lessdef : Proof. split; intros. destruct v; try discriminate; constructor. - unfold valToValue in H. inversion H. unfold intToValue + unfold valToValue in H. inversion H. + Admitted. (* Need to eventually move posToValue 32 to posToValueAuto, as that has this proof. *) Lemma assumption_32bit : @@ -159,11 +160,12 @@ Section CORRECTNESS. Let tge : HTL.genv := HTL.genv_empty. Lemma eval_correct : - forall sp op rs args m v e assoc f, + forall sp op rs args m v v' e assoc f, Op.eval_operation ge sp op (List.map (fun r : BinNums.positive => Registers.Regmap.get r rs) args) m = Some v -> tr_op op args e -> - Verilog.expr_runp f assoc e (valToValue v). + valToValue v = Some v' -> + Verilog.expr_runp f assoc e v'. Admitted. Lemma eval_cond_correct : diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index 39fe3d2..e3b8159 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -17,46 +17,67 @@ *) From coqup Require Import Coquplib Value. -From Coq Require Import FSets.FMapPositive. - From compcert Require Import Maps. -Search Maps.PTree.get not List.In. - Definition reg := positive. -Module AssocMap := PositiveMap. +Module AssocMap := Maps.PTree. Module AssocMapExt. Import AssocMap. - Hint Resolve elements_3w : assocmap. - Hint Resolve elements_correct : assocmap. - Hint Resolve gss : assocmap. - Hint Resolve gso : assocmap. + Hint Resolve elements_correct elements_complete + elements_keys_norepet : assocmap. + Hint Resolve gso gss : assocmap. Section Operations. - Variable elt : Type. + Variable A : Type. - Definition find_default (a : elt) (k : reg) (m : t elt) : elt := - match find k m with + Definition get_default (a : A) (k : reg) (m : t A) : A := + match get k m with | None => a | Some v => v end. - Definition merge (am bm : t elt) : t elt := - fold_right (fun p a => add (fst p) (snd p) a) bm (elements am). + 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 + end. + + Lemma merge_base : + forall am, + merge (empty A) am = am. + Proof. auto. Qed. + Hint Resolve merge_base : assocmap. + + Lemma merge_base2 : + forall am, + merge am (empty A) = am. + Proof. + unfold merge. + destruct am; trivial. + destruct o; trivial. + Qed. + Hint Resolve merge_base2 : assocmap. + + Lemma merge_assoc : + forall am bm k v, + (merge (set k v am) bm) = (set k v (merge am bm)). + Proof. Admitted. + + Definition merge_fold (am bm : t A) : t A := + fold_right (fun p a => set (fst p) (snd p) a) bm (elements am). Lemma add_assoc : - forall k v l bm, + forall (k : elt) (v : A) l bm, List.In (k, v) l -> - SetoidList.NoDupA (@eq_key elt) l -> - @find elt k (fold_right (fun p a => add (fst p) (snd p) a) bm l) = Some v. + list_norepet (List.map fst l) -> + @get A k (fold_right (fun p a => set (fst p) (snd p) a) bm l) = Some v. Proof. - Hint Resolve SetoidList.InA_alt : assocmap. - Hint Extern 1 (exists _, _) => apply list_in_map_inv : assocmap. - induction l; intros. - contradiction. - destruct a as [k' v']. @@ -64,9 +85,7 @@ Module AssocMapExt. + inversion H. inversion H1. inversion H0. subst. simpl. auto with assocmap. - subst. inversion H0. subst. apply in_map with (f := fst) in H1. simpl in *. - unfold not in H4. exfalso. apply H4. apply SetoidList.InA_alt. - auto with assocmap. + inversion H0; subst. apply in_map with (f:=fst) in H1. contradiction. + inversion H. inversion H1. inversion H0. subst. simpl. rewrite gso; try assumption. apply IHl. contradiction. contradiction. @@ -76,9 +95,9 @@ Module AssocMapExt. Lemma not_in_assoc : forall k v l bm, - ~ List.In k (List.map (@fst key elt) l) -> - @find elt k bm = Some v -> - @find elt k (fold_right (fun p a => add (fst p) (snd p) a) bm l) = Some v. + ~ List.In k (List.map (@fst elt A) l) -> + @get A k bm = Some v -> + get k (fold_right (fun p a => set (fst p) (snd p) a) bm l) = Some v. Proof. induction l; intros. - assumption. @@ -91,8 +110,8 @@ Module AssocMapExt. Lemma elements_iff : forall am k, - (exists v, find k am = Some v) <-> - List.In k (List.map (@fst _ elt) (elements am)). + (exists v, get k am = Some v) <-> + List.In k (List.map (@fst _ A) (elements am)). Proof. split; intros. destruct H. apply elements_correct in H. apply in_map with (f := fst) in H. apply H. @@ -104,41 +123,41 @@ Module AssocMapExt. Lemma elements_correct' : forall am k, - ~ (exists v, find k am = Some v) <-> - ~ List.In k (List.map (@fst _ elt) (elements am)). + ~ (exists v, get k am = Some v) <-> + ~ List.In k (List.map (@fst _ A) (elements am)). Proof. auto using not_iff_compat with assocmap. Qed. Hint Resolve elements_correct' : assocmap. Lemma elements_correct_none : forall am k, - find k am = None -> - ~ List.In k (List.map (@fst _ elt) (elements am)). + get k am = None -> + ~ List.In k (List.map (@fst _ A) (elements am)). Proof. intros. apply elements_correct'. unfold not. intros. destruct H0. rewrite H in H0. discriminate. Qed. Hint Resolve elements_correct_none : assocmap. - Lemma merge_add : + Lemma merge_fold_add : forall k v am bm, - find k am = Some v -> - find k (merge am bm) = Some v. - Proof. unfold merge. auto with assocmap. Qed. - Hint Resolve merge_add : assocmap. + am ! k = Some v -> + (merge_fold am bm) ! k = Some v. + Proof. unfold merge_fold; auto with assocmap. Qed. + Hint Resolve merge_fold_add : assocmap. - Lemma merge_not_in : + Lemma merge_fold_not_in : forall k v am bm, - find k am = None -> - find k bm = Some v -> - find k (merge am bm) = Some v. + get k am = None -> + get k bm = Some v -> + get k (merge_fold am bm) = Some v. Proof. intros. apply not_in_assoc; auto with assocmap. Qed. - Hint Resolve merge_not_in : assocmap. + Hint Resolve merge_fold_not_in : assocmap. - Lemma merge_base : + Lemma merge_fold_base : forall am, - merge (empty elt) am = am. + merge_fold (empty A) am = am. Proof. auto. Qed. - Hint Resolve merge_base : assocmap. + Hint Resolve merge_fold_base : assocmap. End Operations. @@ -148,14 +167,14 @@ Import AssocMapExt. Definition assocmap := AssocMap.t value. Definition find_assocmap (n : nat) : reg -> assocmap -> value := - find_default value (ZToValue n 0). + get_default value (ZToValue n 0). Definition empty_assocmap : assocmap := AssocMap.empty value. Definition merge_assocmap : assocmap -> assocmap -> assocmap := merge value. Module AssocMapNotation. - Notation "a ! b" := (AssocMap.find b a) (at level 1). + Notation "a ! b" := (AssocMap.get b a) (at level 1). Notation "a # ( b , c )" := (find_assocmap c b a) (at level 1). Notation "a # b" := (find_assocmap 32 b a) (at level 1). Notation "a ## b" := (List.map (fun c => find_assocmap 32 c a) b) (at level 1). -- cgit From d7ff396015312585336d80c7b4608ac136b7aa0c Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 29 May 2020 15:35:26 +0100 Subject: New and improved Assocmap --- src/verilog/AssocMap.v | 39 ++++++++++++++++++++++++++++++++------- 1 file changed, 32 insertions(+), 7 deletions(-) diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index e3b8159..88ad504 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -48,13 +48,13 @@ Module AssocMapExt. | _, _ => m1 end. - Lemma merge_base : + Lemma merge_base_1 : forall am, merge (empty A) am = am. Proof. auto. Qed. - Hint Resolve merge_base : assocmap. + Hint Resolve merge_base_1 : assocmap. - Lemma merge_base2 : + Lemma merge_base_2 : forall am, merge am (empty A) = am. Proof. @@ -62,12 +62,37 @@ Module AssocMapExt. destruct am; trivial. destruct o; trivial. Qed. - Hint Resolve merge_base2 : assocmap. + Hint Resolve merge_base_2 : assocmap. - Lemma merge_assoc : + 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. + Qed. + Hint Resolve merge_add_assoc : assocmap. + + Lemma merge_correct_1 : forall am bm k v, - (merge (set k v am) bm) = (set k v (merge am bm)). - Proof. Admitted. + 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. + Qed. + Hint Resolve merge_correct_1 : assocmap. + + 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. + Qed. + Hint Resolve merge_correct_2 : assocmap. Definition merge_fold (am bm : t A) : t A := fold_right (fun p a => set (fst p) (snd p) a) bm (elements am). -- cgit From 3f3623f533033aca29fc7c5a05d2dad716133811 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 29 May 2020 15:40:05 +0100 Subject: Fix indentation --- src/verilog/AssocMap.v | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index 88ad504..43f9065 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -50,23 +50,23 @@ Module AssocMapExt. Lemma merge_base_1 : forall am, - merge (empty A) am = am. + merge (empty A) am = am. Proof. auto. Qed. Hint Resolve merge_base_1 : assocmap. Lemma merge_base_2 : forall am, - merge am (empty A) = am. + merge am (empty A) = am. Proof. unfold merge. destruct am; trivial. destruct o; trivial. - Qed. + Qed. Hint Resolve merge_base_2 : assocmap. Lemma merge_add_assoc : forall r am am' v, - merge (set r v am) am' = set r v (merge am am'). + 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. @@ -128,7 +128,7 @@ Module AssocMapExt. - assumption. - destruct a as [k' v']. destruct (peq k k'); subst; - simpl in *; apply Decidable.not_or in H; destruct H. contradiction. + simpl in *; apply Decidable.not_or in H; destruct H. contradiction. rewrite AssocMap.gso; auto. Qed. Hint Resolve not_in_assoc : assocmap. @@ -180,7 +180,7 @@ Module AssocMapExt. Lemma merge_fold_base : forall am, - merge_fold (empty A) am = am. + merge_fold (empty A) am = am. Proof. auto. Qed. Hint Resolve merge_fold_base : assocmap. -- cgit From 6c4e2ba64c17952b432a0f01122c065cae24dacf Mon Sep 17 00:00:00 2001 From: James Pollard Date: Fri, 29 May 2020 18:04:28 +0100 Subject: Improve automation in HTLgenspec. --- src/translation/HTLgenspec.v | 9 +++------ 1 file changed, 3 insertions(+), 6 deletions(-) diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index b9d483c..9c08959 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -350,13 +350,10 @@ Proof. econstructor; simpl; trivial. intros. inv_incr. - assert (st_controllogic s8 = st_controllogic s2). - { rewrite <- H5. rewrite <- H6. rewrite <- H7. trivial. } + assert (st_controllogic s8 = st_controllogic s2) by congruence. rewrite <- H10. - assert (st_datapath s8 = st_datapath s2). - { rewrite <- EQ4. rewrite <- EQ3. rewrite <- EQ2. trivial. } - assert (st_st s5 = st_st s2). - { rewrite H10. rewrite <- H50. trivial. } + assert (st_datapath s8 = st_datapath s2) by congruence. + assert (st_st s5 = st_st s2) by congruence. rewrite H80. rewrite H81. rewrite H82. eauto with htlspec. Qed. -- cgit From e06577fe952a3c268520b280b020bb2bff252529 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 29 May 2020 18:14:59 +0100 Subject: Fix compilation moving to PTree --- src/translation/HTLgen.v | 24 ++++++++++++------------ src/translation/HTLgenproof.v | 31 ++++++++++++++++++++----------- src/translation/HTLgenspec.v | 4 ++-- src/verilog/AssocMap.v | 4 ++++ src/verilog/HTL.v | 7 +++---- src/verilog/Verilog.v | 14 +++++++------- 6 files changed, 48 insertions(+), 36 deletions(-) diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v index 26dc630..4564237 100644 --- a/src/translation/HTLgen.v +++ b/src/translation/HTLgen.v @@ -120,8 +120,8 @@ Lemma add_instr_state_incr : s.(st_st) s.(st_freshreg) (st_freshstate s) - (AssocMap.add n st s.(st_datapath)) - (AssocMap.add n (state_goto s.(st_st) n') s.(st_controllogic))). + (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); @@ -136,8 +136,8 @@ Definition add_instr (n : node) (n' : node) (st : stmnt) : mon unit := s.(st_st) s.(st_freshreg) (st_freshstate s) - (AssocMap.add n st s.(st_datapath)) - (AssocMap.add n (state_goto s.(st_st) n') s.(st_controllogic))) + (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 STM TRANS) | _, _ => Error (Errors.msg "HTL.add_instr") end. @@ -151,8 +151,8 @@ Lemma add_instr_skip_state_incr : s.(st_st) s.(st_freshreg) (st_freshstate s) - (AssocMap.add n st s.(st_datapath)) - (AssocMap.add n Vskip s.(st_controllogic))). + (AssocMap.set n st s.(st_datapath)) + (AssocMap.set n Vskip s.(st_controllogic))). Proof. constructor; intros; try (simpl; destruct (peq n n0); subst); @@ -167,8 +167,8 @@ Definition add_instr_skip (n : node) (st : stmnt) : mon unit := s.(st_st) s.(st_freshreg) (st_freshstate s) - (AssocMap.add n st s.(st_datapath)) - (AssocMap.add n Vskip s.(st_controllogic))) + (AssocMap.set n st s.(st_datapath)) + (AssocMap.set n Vskip s.(st_controllogic))) (add_instr_skip_state_incr s n st STM TRANS) | _, _ => Error (Errors.msg "HTL.add_instr") end. @@ -274,8 +274,8 @@ Lemma add_branch_instr_state_incr: s.(st_st) (st_freshreg s) (st_freshstate s) - (AssocMap.add n Vskip (st_datapath s)) - (AssocMap.add n (state_cond s.(st_st) e n1 n2) (st_controllogic s))). + (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); @@ -290,8 +290,8 @@ Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit := s.(st_st) (st_freshreg s) (st_freshstate s) - (AssocMap.add n Vskip (st_datapath s)) - (AssocMap.add n (state_cond s.(st_st) e n1 n2) (st_controllogic s))) + (AssocMap.set n Vskip (st_datapath s)) + (AssocMap.set n (state_cond s.(st_st) e n1 n2) (st_controllogic s))) (add_branch_instr_state_incr s e n n1 n2 NSTM NTRANS) | _, _ => Error (Errors.msg "Veriloggen: add_branch_instr") end. diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v index b7b9d27..2d2129c 100644 --- a/src/translation/HTLgenproof.v +++ b/src/translation/HTLgenproof.v @@ -26,6 +26,8 @@ Hint Resolve Smallstep.forward_simulation_plus : htlproof. Hint Resolve AssocMap.gss : htlproof. Hint Resolve AssocMap.gso : htlproof. +Hint Unfold find_assocmap AssocMapExt.get_default : htlproof. + Inductive match_assocmaps : RTL.function -> RTL.regset -> assocmap -> Prop := match_assocmap : forall f rs am, (forall r, Ple r (RTL.max_reg_function f) -> @@ -61,37 +63,39 @@ Inductive match_program : RTL.program -> HTL.module -> Prop := match_program p m. Hint Constructors match_program : htlproof. -Lemma regs_lessdef_regs : +Lemma regs_lessdef_add_greater : forall f rs1 rs2 n v, Plt (RTL.max_reg_function f) n -> match_assocmaps f rs1 rs2 -> - match_assocmaps f rs1 (AssocMap.add n v rs2). + match_assocmaps f rs1 (AssocMap.set n v rs2). Proof. inversion 2; subst. intros. constructor. - intros. unfold find_assocmap. unfold AssocMapExt.find_default. + intros. unfold find_assocmap. unfold AssocMapExt.get_default. rewrite AssocMap.gso. eauto. apply Pos.le_lt_trans with _ _ n in H2. unfold not. intros. subst. eapply Pos.lt_irrefl. eassumption. assumption. Qed. +Hint Resolve regs_lessdef_add_greater : htlproof. -Lemma regs_add_match : +Lemma regs_lessdef_add_match : forall f rs am r v v', val_value_lessdef v v' -> match_assocmaps f rs am -> - match_assocmaps f (Registers.Regmap.set r v rs) (AssocMap.add r v' am). + match_assocmaps f (Registers.Regmap.set r v rs) (AssocMap.set r v' am). Proof. inversion 2; subst. constructor. intros. destruct (peq r0 r); subst. rewrite Registers.Regmap.gss. - unfold find_assocmap. unfold AssocMapExt.find_default. + unfold find_assocmap. unfold AssocMapExt.get_default. rewrite AssocMap.gss. assumption. rewrite Registers.Regmap.gso; try assumption. - unfold find_assocmap. unfold AssocMapExt.find_default. + unfold find_assocmap. unfold AssocMapExt.get_default. rewrite AssocMap.gso; eauto. Qed. +Hint Resolve regs_lessdef_add_match : htlproof. Lemma valToValue_lessdef : forall v v', @@ -198,6 +202,11 @@ Section CORRECTNESS. exists assoc', HTL.step tge (HTL.State m st assoc) Events.E0 (HTL.State m st assoc'). + Lemma greater_than_max_func : + forall f st, + Plt (RTL.max_reg_function f) st. + Proof. Admitted. + Theorem transl_step_correct: forall (S1 : RTL.state) t S2, RTL.step ge S1 t S2 -> @@ -218,14 +227,14 @@ Section CORRECTNESS. econstructor. (* prove stval *) - unfold merge_assocmap. apply AssocMapExt.merge_add. simpl. apply AssocMap.gss. + unfold_merge. simpl. apply AssocMap.gss. (* prove match_state *) rewrite assumption_32bit. constructor; auto. - unfold merge_assocmap. rewrite AssocMapExt.merge_add. simpl. apply AssocMap.gss. - apply regs_lessdef_regs. assumption. - unfold state_st_wf. inversion 1. subst. apply AssocMap.gss. + unfold_merge. simpl. apply regs_lessdef_regs. apply greater_than_max_func. + assumption. + unfold state_st_wf. inversion 1. subst. unfold_merge. apply AssocMap.gss. - (* Iop *) econstructor. split. diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v index b9d483c..5a61a9c 100644 --- a/src/translation/HTLgenspec.v +++ b/src/translation/HTLgenspec.v @@ -16,8 +16,8 @@ * along with this program. If not, see . *) -From Coq Require Import FSets.FMapPositive. From compcert Require RTL Op Maps Errors. +From compcert Require Import Maps. From coqup Require Import Coquplib Verilog Value HTL HTLgen AssocMap. Hint Resolve Maps.PTree.elements_keys_norepet : htlspec. @@ -160,7 +160,7 @@ Inductive tr_instr (fin rtrn st : reg) : RTL.instruction -> stmnt -> stmnt -> Pr (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r))) Vskip. Hint Constructors tr_instr : htlspec. -Inductive tr_code (pc : RTL.node) (i : RTL.instruction) (stmnts trans : PositiveMap.t stmnt) +Inductive tr_code (pc : RTL.node) (i : RTL.instruction) (stmnts trans : PTree.t stmnt) (fin rtrn st : reg) : Prop := tr_code_intro : forall s t, diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v index 43f9065..88b13a6 100644 --- a/src/verilog/AssocMap.v +++ b/src/verilog/AssocMap.v @@ -198,6 +198,10 @@ Definition empty_assocmap : assocmap := AssocMap.empty value. Definition merge_assocmap : assocmap -> assocmap -> assocmap := merge value. +Ltac unfold_merge := + unfold merge_assocmap; try (repeat (rewrite merge_add_assoc)); + rewrite AssocMapExt.merge_base_1. + Module AssocMapNotation. Notation "a ! b" := (AssocMap.get b a) (at level 1). Notation "a # ( b , c )" := (find_assocmap c b a) (at level 1). diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v index 36e4434..a21064c 100644 --- a/src/verilog/HTL.v +++ b/src/verilog/HTL.v @@ -20,6 +20,7 @@ From Coq Require Import FSets.FMapPositive. From coqup Require Import Coquplib Value AssocMap. From coqup Require Verilog. From compcert Require Events Globalenvs Smallstep Integers. +From compcert Require Import Maps. Import HexNotationValue. @@ -29,13 +30,11 @@ hardware-like layout that is still similar to the register transfer language instantiations and that we now describe a state machine instead of a control-flow graph. *) -Notation "a ! b" := (PositiveMap.find b a) (at level 1). - Definition reg := positive. Definition node := positive. -Definition datapath := PositiveMap.t Verilog.stmnt. -Definition controllogic := PositiveMap.t Verilog.stmnt. +Definition datapath := PTree.t Verilog.stmnt. +Definition controllogic := PTree.t Verilog.stmnt. Record module: Type := mkmodule { diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v index c1e0a79..ce7ddd9 100644 --- a/src/verilog/Verilog.v +++ b/src/verilog/Verilog.v @@ -171,13 +171,13 @@ Definition posToLit (p : positive) : expr := Coercion Vlit : value >-> expr. Coercion Vvar : reg >-> expr. -Definition fext := PositiveMap.t value. +Definition fext := AssocMap.t value. Definition fextclk := nat -> fext. (** ** 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. @@ -300,7 +300,7 @@ 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 PositiveMap.find r f with + match AssocMap.get r f with | Some v => OK v | _ => OK (ZToValue 1 0) end. @@ -452,7 +452,7 @@ Inductive stmnt_runp: fext -> associations -> stmnt -> associations -> Prop := forall lhs name rhs rhsval assoc assoc' nbassoc f, assign_name lhs = OK name -> expr_runp f assoc rhs rhsval -> - assoc' = (PositiveMap.add name rhsval assoc) -> + assoc' = (AssocMap.set name rhsval assoc) -> stmnt_runp f (mkassociations assoc nbassoc) (Vblock lhs rhs) (mkassociations assoc' nbassoc) @@ -460,7 +460,7 @@ Inductive stmnt_runp: fext -> associations -> stmnt -> associations -> Prop := forall lhs name rhs rhsval assoc nbassoc nbassoc' f, assign_name lhs = OK name -> expr_runp f assoc rhs rhsval -> - nbassoc' = (PositiveMap.add name rhsval nbassoc) -> + nbassoc' = (AssocMap.set name rhsval nbassoc) -> stmnt_runp f (mkassociations assoc nbassoc) (Vnonblock lhs rhs) (mkassociations assoc nbassoc'). @@ -534,8 +534,8 @@ other arguments to the module are also to be supported. *) Definition initial_fextclk (m : module) : fextclk := fun x => match x with - | S O => (AssocMap.add (fst (mod_reset m)) (ZToValue 1 1) empty_assocmap) - | _ => (AssocMap.add (fst (mod_reset m)) (ZToValue 1 0) empty_assocmap) + | S O => (AssocMap.set (fst (mod_reset m)) (ZToValue 1 1) empty_assocmap) + | _ => (AssocMap.set (fst (mod_reset m)) (ZToValue 1 0) empty_assocmap) end. (*Definition module_run (n : nat) (m : module) : res assocmap := -- cgit