aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2020-06-25 09:09:23 +0100
committerYann Herklotz <git@yannherklotz.com>2020-06-25 09:09:23 +0100
commit445aabbcf63e29d68dd0c98dde7f259af0381591 (patch)
treefe77b0b489849e60dca1fc221f7c3a301ec776e5
parent14cd3b8b12e3db17c3842ae9dfdb30ca86b6394c (diff)
downloadvericert-445aabbcf63e29d68dd0c98dde7f259af0381591.tar.gz
vericert-445aabbcf63e29d68dd0c98dde7f259af0381591.zip
Work on Veriloggen proof
-rw-r--r--src/translation/HTLgenproof.v2
-rw-r--r--src/translation/Veriloggenproof.v52
-rw-r--r--src/verilog/Verilog.v7
3 files changed, 55 insertions, 6 deletions
diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v
index f07403d..77bd4ef 100644
--- a/src/translation/HTLgenproof.v
+++ b/src/translation/HTLgenproof.v
@@ -1424,7 +1424,7 @@ Section CORRECTNESS.
assert (Some (AST.Internal x) = Some (AST.Internal m)).
replace (AST.fundef HTL.module) with (HTL.fundef).
rewrite <- H6. setoid_rewrite <- A. trivial.
- trivial. inversion H7.
+ trivial. inv H7. assumption.
Qed.
Hint Resolve transl_initial_states : htlproof.
diff --git a/src/translation/Veriloggenproof.v b/src/translation/Veriloggenproof.v
index 6c58c56..825cb7e 100644
--- a/src/translation/Veriloggenproof.v
+++ b/src/translation/Veriloggenproof.v
@@ -16,16 +16,64 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-From compcert Require Import Smallstep.
+From compcert Require Import Smallstep Linking.
+From coqup Require Import Veriloggen.
From coqup Require HTL Verilog.
+Definition match_prog (prog : HTL.program) (tprog : Verilog.program) :=
+ match_program (fun cu f tf => tf = transl_fundef f) eq prog tprog.
+
+Lemma transf_program_match:
+ forall prog, match_prog prog (transl_program prog).
+Proof.
+ intros. eapply match_transform_program_contextual. auto.
+Qed.
+
+Inductive match_stacks : list HTL.stackframe -> list Verilog.stackframe -> Prop :=
+| match_stack :
+ forall res m pc reg_assoc arr_assoc hstk vstk,
+ match_stacks hstk vstk ->
+ match_stacks (HTL.Stackframe res m pc reg_assoc arr_assoc :: hstk)
+ (Verilog.Stackframe res (transl_module m) pc
+ reg_assoc arr_assoc :: vstk)
+| match_stack_nil : match_stacks nil nil.
+
+Inductive match_states : HTL.state -> Verilog.state -> Prop :=
+| match_state :
+ forall m st reg_assoc arr_assoc hstk vstk,
+ match_stacks hstk vstk ->
+ match_states (HTL.State hstk m st reg_assoc arr_assoc)
+ (Verilog.State vstk (transl_module m) st reg_assoc arr_assoc)
+| match_returnstate :
+ forall v hstk vstk,
+ match_stacks hstk vstk ->
+ match_states (HTL.Returnstate hstk v) (Verilog.Returnstate vstk v)
+| match_initial_call :
+ forall m,
+ match_states (HTL.Callstate nil m nil) (Verilog.Callstate nil (transl_module m) nil).
+
Section CORRECTNESS.
Variable prog: HTL.program.
Variable tprog: Verilog.program.
+ Hypothesis TRANSL : match_prog prog tprog.
+
+ Let ge : HTL.genv := Globalenvs.Genv.globalenv prog.
+ Let tge : Verilog.genv := Globalenvs.Genv.globalenv tprog.
+
+ Theorem transl_step_correct:
+ forall (S1 : HTL.state) t S2,
+ HTL.step ge S1 t S2 ->
+ forall (R1 : Verilog.state),
+ match_states S1 R1 ->
+ exists R2, Smallstep.plus Verilog.step tge R1 t R2 /\ match_states S2 R2.
+ Proof.
+ induction 1; intros R1 MSTATE.
+
Theorem transf_program_correct:
forward_simulation (HTL.semantics prog) (Verilog.semantics tprog).
- Admitted.
+
End CORRECTNESS.
+
diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v
index 7d5e3c0..9c05fc9 100644
--- a/src/verilog/Verilog.v
+++ b/src/verilog/Verilog.v
@@ -279,7 +279,8 @@ Inductive stackframe : Type :=
forall (res : reg)
(m : module)
(pc : node)
- (assoc : assocmap),
+ (reg_assoc : assocmap_reg)
+ (arr_assoc : assocmap_arr),
stackframe.
Inductive state : Type :=
@@ -735,9 +736,9 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop :=
(init_params args m.(mod_args)))
(empty_stack m))
| step_return :
- forall g m asr i r sf pc mst,
+ forall g m asr i r sf pc mst asa,
mst = mod_st m ->
- step g (Returnstate (Stackframe r m pc asr :: sf) i) Events.E0
+ step g (Returnstate (Stackframe r m pc asr asa :: sf) i) Events.E0
(State sf m pc ((asr # mst <- (posToValue 32 pc)) # r <- i)
(empty_stack m)).
Hint Constructors step : verilog.