aboutsummaryrefslogtreecommitdiffstats
path: root/src/verilog/Test.v
blob: 5fd6d07f040d424d8b43cc7549a1a4997ad5ef6a (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
(*
 * CoqUp: Verified high-level synthesis.
 * Copyright (C) 2020 Yann Herklotz <yann@yannherklotz.com>
 *
 * 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 <https://www.gnu.org/licenses/>.
 *)

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 :=
  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] |}.

Search (_ <> _ -> _ = _).

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 :
  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. 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.