aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorJames Pollard <james@pollard.dev>2020-07-07 13:16:47 +0100
committerJames Pollard <james@pollard.dev>2020-07-07 13:16:47 +0100
commit5baa025e8d26b91ac99f983bce44af5025e20b14 (patch)
treefaa3fadc4483127929f629031e711f61f4dce411
parent7f7ad4a7a60c9e914f1e05cb37d3166ba11bacc6 (diff)
parent6c9fad2cf278df9dccbc7a7167f7774acbed280b (diff)
downloadvericert-5baa025e8d26b91ac99f983bce44af5025e20b14.tar.gz
vericert-5baa025e8d26b91ac99f983bce44af5025e20b14.zip
Merge branch 'byte-addressing' into dev-nadesh
-rw-r--r--driver/CoqupDriver.ml4
-rw-r--r--src/Compiler.v136
-rw-r--r--src/CoqupClflags.ml1
-rw-r--r--src/common/Coquplib.v76
-rw-r--r--src/common/IntegerExtra.v167
-rw-r--r--src/common/Monad.v4
-rw-r--r--src/common/ZExtra.v15
-rw-r--r--src/extraction/Extraction.v5
-rw-r--r--src/translation/HTLgen.v210
-rw-r--r--src/translation/HTLgenproof.v3558
-rw-r--r--src/translation/HTLgenspec.v268
-rw-r--r--src/translation/Veriloggen.v36
-rw-r--r--src/translation/Veriloggenproof.v303
-rw-r--r--src/translation/Veriloggenspec.v17
-rw-r--r--src/verilog/Array.v69
-rw-r--r--src/verilog/AssocMap.v4
-rw-r--r--src/verilog/HTL.v34
-rw-r--r--src/verilog/PrintHTL.ml81
-rw-r--r--src/verilog/PrintVerilog.ml22
-rw-r--r--src/verilog/PrintVerilog.mli6
-rw-r--r--src/verilog/Value.v107
-rw-r--r--src/verilog/ValueInt.v162
-rw-r--r--src/verilog/Verilog.v341
-rw-r--r--test/array.c7
-rw-r--r--test/matrix.c16
25 files changed, 3436 insertions, 2213 deletions
diff --git a/driver/CoqupDriver.ml b/driver/CoqupDriver.ml
index afbe4d0..0b64f48 100644
--- a/driver/CoqupDriver.ml
+++ b/driver/CoqupDriver.ml
@@ -65,6 +65,7 @@ let compile_c_file sourcename ifile ofile =
set_dest Coqup.PrintClight.destination option_dclight ".light.c";
set_dest Coqup.PrintCminor.destination option_dcminor ".cm";
set_dest Coqup.PrintRTL.destination option_drtl ".rtl";
+ set_dest Coqup.PrintHTL.destination option_dhtl ".htl";
set_dest Coqup.Regalloc.destination_alloctrace option_dalloctrace ".alloctrace";
set_dest Coqup.PrintLTL.destination option_dltl ".ltl";
set_dest Coqup.PrintMach.destination option_dmach ".mach";
@@ -253,6 +254,7 @@ Code generation options: (use -fno-<opt> to turn off -f<opt>)
-dclight Save generated Clight in <file>.light.c
-dcminor Save generated Cminor in <file>.cm
-drtl Save RTL at various optimization points in <file>.rtl.<n>
+ -dhtl Save HTL before Verilog generation <file>.htl
-dltl Save LTL after register allocation in <file>.ltl
-dmach Save generated Mach code in <file>.mach
-dasm Save generated assembly in <file>.s
@@ -377,6 +379,7 @@ let cmdline_actions =
Exact "-dclight", Set option_dclight;
Exact "-dcminor", Set option_dcminor;
Exact "-drtl", Set option_drtl;
+ Exact "-dhtl", Set option_dhtl;
Exact "-dltl", Set option_dltl;
Exact "-dalloctrace", Set option_dalloctrace;
Exact "-dmach", Set option_dmach;
@@ -388,6 +391,7 @@ let cmdline_actions =
option_dclight := true;
option_dcminor := true;
option_drtl := true;
+ option_dhtl := true;
option_dltl := true;
option_dalloctrace := true;
option_dmach := true;
diff --git a/src/Compiler.v b/src/Compiler.v
index f9fe686..26e2f1f 100644
--- a/src/Compiler.v
+++ b/src/Compiler.v
@@ -16,6 +16,8 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
+From coqup Require Import HTLgenproof.
+
From compcert.common Require Import
Errors
Linking.
@@ -48,9 +50,13 @@ From compcert.driver Require
From coqup Require
Verilog
Veriloggen
+ Veriloggenproof
HTLgen.
+From compcert Require Import Smallstep.
+
Parameter print_RTL: Z -> RTL.program -> unit.
+Parameter print_HTL: HTL.program -> unit.
Definition print {A: Type} (printer: A -> unit) (prog: A) : A :=
let unused := printer prog in prog.
@@ -86,18 +92,10 @@ Definition transf_backend (r : RTL.program) : res Verilog.program :=
@@@ Unusedglob.transform_program
@@ print (print_RTL 1)
@@@ HTLgen.transl_program
+ @@ print print_HTL
@@ Veriloggen.transl_program.
-(* Unoptimised below: *)
-(*Definition transf_backend (r : RTL.program) : res Verilog.program :=
- OK r
- @@@ Inlining.transf_program
- @@ print (print_RTL 1)
- @@@ HTLgen.transl_program
- @@ Veriloggen.transl_program.
-*)
-
-Definition transf_frontend (p: Csyntax.program) : res RTL.program :=
+Definition transf_hls (p : Csyntax.program) : res Verilog.program :=
OK p
@@@ SimplExpr.transl_program
@@@ SimplLocals.transf_program
@@ -105,11 +103,6 @@ Definition transf_frontend (p: Csyntax.program) : res RTL.program :=
@@@ Cminorgen.transl_program
@@@ Selection.sel_program
@@@ RTLgen.transl_program
- @@ print (print_RTL 0).
-
-Definition transf_hls (p : Csyntax.program) : res Verilog.program :=
- OK p
- @@@ transf_frontend
@@@ transf_backend.
Local Open Scope linking_scope.
@@ -121,25 +114,31 @@ Definition CompCert's_passes :=
::: mkpass Cminorgenproof.match_prog
::: mkpass Selectionproof.match_prog
::: mkpass RTLgenproof.match_prog
+ ::: mkpass Inliningproof.match_prog
+ ::: (@mkpass _ _ HTLgenproof.match_prog (HTLgenproof.TransfHTLLink HTLgen.transl_program))
+ ::: mkpass Veriloggenproof.match_prog
::: pass_nil _.
-Definition match_prog: Csyntax.program -> RTL.program -> Prop :=
+Definition match_prog: Csyntax.program -> Verilog.program -> Prop :=
pass_match (compose_passes CompCert's_passes).
-Theorem transf_frontend_match:
+Theorem transf_hls_match:
forall p tp,
- transf_frontend p = OK tp ->
+ transf_hls p = OK tp ->
match_prog p tp.
Proof.
intros p tp T.
- unfold transf_frontend in T. simpl in T.
+ unfold transf_hls in T. simpl in T.
destruct (SimplExpr.transl_program p) as [p1|e] eqn:P1; simpl in T; try discriminate.
destruct (SimplLocals.transf_program p1) as [p2|e] eqn:P2; simpl in T; try discriminate.
destruct (Cshmgen.transl_program p2) as [p3|e] eqn:P3; simpl in T; try discriminate.
destruct (Cminorgen.transl_program p3) as [p4|e] eqn:P4; simpl in T; try discriminate.
destruct (Selection.sel_program p4) as [p5|e] eqn:P5; simpl in T; try discriminate.
- rewrite ! compose_print_identity in T.
destruct (RTLgen.transl_program p5) as [p6|e] eqn:P6; simpl in T; try discriminate.
+ unfold transf_backend in T. simpl in T. rewrite ! compose_print_identity in T.
+ destruct (Inlining.transf_program p6) as [p7|e] eqn:P7; simpl in T; try discriminate.
+ destruct (HTLgen.transl_program p7) as [p8|e] eqn:P8; simpl in T; try discriminate.
+ set (p9 := Veriloggen.transl_program p8) in *.
unfold match_prog; simpl.
exists p1; split. apply SimplExprproof.transf_program_match; auto.
exists p2; split. apply SimplLocalsproof.match_transf_program; auto.
@@ -147,5 +146,100 @@ Proof.
exists p4; split. apply Cminorgenproof.transf_program_match; auto.
exists p5; split. apply Selectionproof.transf_program_match; auto.
exists p6; split. apply RTLgenproof.transf_program_match; auto.
- inversion T. reflexivity.
+ exists p7; split. apply Inliningproof.transf_program_match; auto.
+ exists p8; split. apply HTLgenproof.transf_program_match; auto.
+ exists p9; split. apply Veriloggenproof.transf_program_match; auto.
+ inv T. reflexivity.
+Qed.
+
+Remark forward_simulation_identity:
+ forall sem, forward_simulation sem sem.
+Proof.
+ intros. apply forward_simulation_step with (fun s1 s2 => s2 = s1); intros.
+- auto.
+- exists s1; auto.
+- subst s2; auto.
+- subst s2. exists s1'; auto.
+Qed.
+
+Theorem cstrategy_semantic_preservation:
+ forall p tp,
+ match_prog p tp ->
+ forward_simulation (Cstrategy.semantics p) (Verilog.semantics tp)
+ /\ backward_simulation (atomic (Cstrategy.semantics p)) (Verilog.semantics tp).
+Proof.
+ intros p tp M. unfold match_prog, pass_match in M; simpl in M.
+Ltac DestructM :=
+ match goal with
+ [ H: exists p, _ /\ _ |- _ ] =>
+ let p := fresh "p" in let M := fresh "M" in let MM := fresh "MM" in
+ destruct H as (p & M & MM); clear H
+ end.
+ repeat DestructM. subst tp.
+ assert (F: forward_simulation (Cstrategy.semantics p) (Verilog.semantics p9)).
+ {
+ eapply compose_forward_simulations.
+ eapply SimplExprproof.transl_program_correct; eassumption.
+ eapply compose_forward_simulations.
+ eapply SimplLocalsproof.transf_program_correct; eassumption.
+ eapply compose_forward_simulations.
+ eapply Cshmgenproof.transl_program_correct; eassumption.
+ eapply compose_forward_simulations.
+ eapply Cminorgenproof.transl_program_correct; eassumption.
+ eapply compose_forward_simulations.
+ eapply Selectionproof.transf_program_correct; eassumption.
+ eapply compose_forward_simulations.
+ eapply RTLgenproof.transf_program_correct; eassumption.
+ eapply compose_forward_simulations.
+ eapply Inliningproof.transf_program_correct; eassumption.
+ eapply compose_forward_simulations.
+ eapply HTLgenproof.transf_program_correct. eassumption.
+ eapply Veriloggenproof.transf_program_correct; eassumption.
+ }
+ split. auto.
+ apply forward_to_backward_simulation.
+ apply factor_forward_simulation. auto. eapply sd_traces. eapply Verilog.semantics_determinate.
+ apply atomic_receptive. apply Cstrategy.semantics_strongly_receptive.
+ apply Verilog.semantics_determinate.
+Qed.
+
+Theorem c_semantic_preservation:
+ forall p tp,
+ match_prog p tp ->
+ backward_simulation (Csem.semantics p) (Verilog.semantics tp).
+Proof.
+ intros.
+ apply compose_backward_simulation with (atomic (Cstrategy.semantics p)).
+ eapply sd_traces; eapply Verilog.semantics_determinate.
+ apply factor_backward_simulation.
+ apply Cstrategy.strategy_simulation.
+ apply Csem.semantics_single_events.
+ eapply ssr_well_behaved; eapply Cstrategy.semantics_strongly_receptive.
+ exact (proj2 (cstrategy_semantic_preservation _ _ H)).
+Qed.
+
+Theorem transf_c_program_correct:
+ forall p tp,
+ transf_hls p = OK tp ->
+ backward_simulation (Csem.semantics p) (Verilog.semantics tp).
+Proof.
+ intros. apply c_semantic_preservation. apply transf_hls_match; auto.
+Qed.
+
+Theorem separate_transf_c_program_correct:
+ forall c_units verilog_units c_program,
+ nlist_forall2 (fun cu tcu => transf_hls cu = OK tcu) c_units verilog_units ->
+ link_list c_units = Some c_program ->
+ exists verilog_program,
+ link_list verilog_units = Some verilog_program
+ /\ backward_simulation (Csem.semantics c_program) (Verilog.semantics verilog_program).
+Proof.
+ intros.
+ assert (nlist_forall2 match_prog c_units verilog_units).
+ { eapply nlist_forall2_imply. eauto. simpl; intros. apply transf_hls_match; auto. }
+ assert (exists verilog_program, link_list verilog_units = Some verilog_program
+ /\ match_prog c_program verilog_program).
+ { eapply link_list_compose_passes; eauto. }
+ destruct H2 as (verilog_program & P & Q).
+ exists verilog_program; split; auto. apply c_semantic_preservation; auto.
Qed.
diff --git a/src/CoqupClflags.ml b/src/CoqupClflags.ml
index 83dd31d..5b40ce6 100644
--- a/src/CoqupClflags.ml
+++ b/src/CoqupClflags.ml
@@ -3,3 +3,4 @@ let option_simulate = ref false
let option_hls = ref true
let option_debug_hls = ref false
let option_initial = ref false
+let option_dhtl = ref false
diff --git a/src/common/Coquplib.v b/src/common/Coquplib.v
index 5de1e7c..469eddc 100644
--- a/src/common/Coquplib.v
+++ b/src/common/Coquplib.v
@@ -23,12 +23,37 @@ From Coq Require Export
List
Bool.
+Require Import Lia.
+
From coqup Require Import Show.
(* Depend on CompCert for the basic library, as they declare and prove some
useful theorems. *)
From compcert.lib Require Export Coqlib.
-From compcert Require Integers.
+From compcert Require Import Integers.
+
+Local Open Scope Z_scope.
+
+(* This tactic due to Clement Pit-Claudel with some minor additions by JDP to
+ allow the result to be named: https://pit-claudel.fr/clement/MSc/#org96a1b5f *)
+Inductive Learnt {A: Type} (a: A) :=
+ | AlreadyKnown : Learnt a.
+
+Ltac learn_tac fact name :=
+ lazymatch goal with
+ | [ H: Learnt fact |- _ ] =>
+ fail 0 "fact" fact "has already been learnt"
+ | _ => let type := type of fact in
+ lazymatch goal with
+ | [ H: @Learnt type _ |- _ ] =>
+ fail 0 "fact" fact "of type" type "was already learnt through" H
+ | _ => let learnt := fresh "Learn" in
+ pose proof (AlreadyKnown fact) as learnt; pose proof fact as name
+ end
+ end.
+
+Tactic Notation "learn" constr(fact) := let name := fresh "H" in learn_tac fact name.
+Tactic Notation "learn" constr(fact) "as" simple_intropattern(name) := learn_tac fact name.
Ltac unfold_rec c := unfold c; fold c.
@@ -44,13 +69,24 @@ Ltac solve_by_invert := solve_by_inverts 1.
Ltac invert x := inversion x; subst; clear x.
-Ltac clear_obvious :=
+Ltac destruct_match :=
+ match goal with | [ |- context[match ?x with | _ => _ end ] ] => destruct x end.
+
+Ltac nicify_hypotheses :=
repeat match goal with
| [ H : ex _ |- _ ] => invert H
- | [ H : ?C _ = ?C _ |- _ ] => invert H
+ | [ H : Some _ = Some _ |- _ ] => invert H
+ | [ H : ?x = ?x |- _ ] => clear H
| [ H : _ /\ _ |- _ ] => invert H
end.
+Ltac nicify_goals :=
+ repeat match goal with
+ | [ |- _ /\ _ ] => split
+ | [ |- Some _ = Some _ ] => f_equal
+ | [ |- S _ = S _ ] => f_equal
+ end.
+
Ltac kill_bools :=
repeat match goal with
| [ H : _ && _ = true |- _ ] => apply andb_prop in H
@@ -117,16 +153,38 @@ Ltac unfold_constants :=
end
end.
-Ltac simplify := unfold_constants; simpl in *;
- repeat (clear_obvious; kill_bools);
- simpl in *; try discriminate.
+Ltac substpp :=
+ repeat match goal with
+ | [ H1 : ?x = Some _, H2 : ?x = Some _ |- _ ] =>
+ let EQ := fresh "EQ" in
+ learn H1 as EQ; rewrite H2 in EQ; invert EQ
+ | _ => idtac
+ end.
-Global Opaque Nat.div.
-Global Opaque Z.mul.
+Ltac simplify := intros; unfold_constants; simpl in *;
+ repeat (nicify_hypotheses; nicify_goals; kill_bools; substpp);
+ simpl in *.
Infix "==nat" := eq_nat_dec (no associativity, at level 50).
Infix "==Z" := Z.eq_dec (no associativity, at level 50).
+Ltac liapp :=
+ repeat match goal with
+ | [ |- (?x | ?y) ] =>
+ match (eval compute in (Z.rem y x ==Z 0)) with
+ | left _ =>
+ let q := (eval compute in (Z.div y x))
+ in exists q; reflexivity
+ | _ => idtac
+ end
+ | _ => idtac
+ end.
+
+Ltac crush := simplify; try discriminate; try congruence; try lia; liapp; try assumption.
+
+Global Opaque Nat.div.
+Global Opaque Z.mul.
+
(* Definition const (A B : Type) (a : A) (b : B) : A := a.
Definition compose (A B C : Type) (f : B -> C) (g : A -> B) (x : A) : C := f (g x). *)
@@ -177,3 +235,5 @@ Definition debug_show {A B : Type} `{Show A} (a : A) (b : B) : B :=
Definition debug_show_msg {A B : Type} `{Show A} (s : string) (a : A) (b : B) : B :=
let unused := debug_print (s ++ show a) in b.
+
+Notation "f $ x" := (f x) (at level 60, right associativity, only parsing).
diff --git a/src/common/IntegerExtra.v b/src/common/IntegerExtra.v
index 7d3156b..8e32c2c 100644
--- a/src/common/IntegerExtra.v
+++ b/src/common/IntegerExtra.v
@@ -27,7 +27,7 @@ Module PtrofsExtra.
rewrite Zmod_mod
| [ _ : _ |- context[(_ mod Ptrofs.modulus) mod m ] ] =>
rewrite <- Zmod_div_mod;
- try (simplify; lia || assumption)
+ try (crush; lia || assumption)
| [ _ : _ |- context[Ptrofs.modulus mod m] ] =>
rewrite Zdivide_mod with (a := Ptrofs.modulus);
@@ -65,27 +65,26 @@ Module PtrofsExtra.
| [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] =>
rewrite <- Zmod_div_mod; try lia; try assumption
| [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed
- end; try (simplify; lia); ptrofs_mod_tac m.
+ end; try crush; ptrofs_mod_tac m.
Qed.
Lemma of_int_mod :
forall x m,
- Int.signed x mod m = 0 ->
- Ptrofs.signed (Ptrofs.of_int x) mod m = 0.
+ Int.unsigned x mod m = 0 ->
+ Ptrofs.unsigned (Ptrofs.of_int x) mod m = 0.
Proof.
intros.
- pose proof (Integers.Ptrofs.agree32_of_int eq_refl x) as A.
- pose proof Ptrofs.agree32_signed.
- apply H0 in A; try reflexivity.
- rewrite A. assumption.
+ unfold Ptrofs.of_int.
+ rewrite Ptrofs.unsigned_repr; crush;
+ apply Int.unsigned_range_2.
Qed.
Lemma mul_mod :
forall x y m,
0 < m ->
(m | Ptrofs.modulus) ->
- Ptrofs.signed x mod m = 0 ->
- Ptrofs.signed y mod m = 0 ->
+ Ptrofs.unsigned x mod m = 0 ->
+ Ptrofs.unsigned y mod m = 0 ->
(Ptrofs.signed (Ptrofs.mul x y)) mod m = 0.
Proof.
intros. unfold Ptrofs.mul.
@@ -95,16 +94,15 @@ Module PtrofsExtra.
| [ _ : _ |- context[if ?x then _ else _] ] => destruct x
| [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] =>
rewrite <- Zmod_div_mod; try lia; try assumption
- | [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed
- end; try (simplify; lia); ptrofs_mod_tac m.
+ end; try(crush; lia); ptrofs_mod_tac m.
Qed.
Lemma add_mod :
forall x y m,
0 < m ->
(m | Ptrofs.modulus) ->
- Ptrofs.signed x mod m = 0 ->
- Ptrofs.signed y mod m = 0 ->
+ Ptrofs.unsigned x mod m = 0 ->
+ Ptrofs.unsigned y mod m = 0 ->
(Ptrofs.unsigned (Ptrofs.add x y)) mod m = 0.
Proof.
intros. unfold Ptrofs.add.
@@ -114,8 +112,7 @@ Module PtrofsExtra.
| [ _ : _ |- context[if ?x then _ else _] ] => destruct x
| [ _ : _ |- context[_ mod Ptrofs.modulus mod m] ] =>
rewrite <- Zmod_div_mod; try lia; try assumption
- | [ _ : _ |- context[Ptrofs.unsigned _] ] => rewrite Ptrofs.unsigned_signed
- end; try (simplify; lia); ptrofs_mod_tac m.
+ end; try (crush; lia); ptrofs_mod_tac m.
Qed.
Lemma mul_divu :
@@ -156,7 +153,7 @@ Module PtrofsExtra.
eapply Z.le_trans.
exact H0.
rewrite Z.mul_comm.
- apply Z.le_mul_diag_r; simplify; lia.
+ apply Z.le_mul_diag_r; crush.
Qed.
Lemma mul_unsigned :
@@ -184,8 +181,25 @@ Module PtrofsExtra.
Qed.
End PtrofsExtra.
-Module IntExtra.
+Ltac ptrofs :=
+ repeat match goal with
+ | [ |- context[Ptrofs.add (Ptrofs.zero) _] ] => setoid_rewrite Ptrofs.add_zero_l
+ | [ H : context[Ptrofs.add (Ptrofs.zero) _] |- _ ] => setoid_rewrite Ptrofs.add_zero_l in H
+
+ | [ |- context[Ptrofs.repr 0] ] => replace (Ptrofs.repr 0) with Ptrofs.zero by reflexivity
+ | [ H : context[Ptrofs.repr 0] |- _ ] =>
+ replace (Ptrofs.repr 0) with Ptrofs.zero in H by reflexivity
+ | [ H: context[Ptrofs.unsigned (Ptrofs.repr (Ptrofs.unsigned _))] |- _ ] =>
+ setoid_rewrite Ptrofs.unsigned_repr in H; [>| apply Ptrofs.unsigned_range_2]
+ | [ |- context[Ptrofs.unsigned (Ptrofs.repr (Ptrofs.unsigned _))] ] =>
+ rewrite Ptrofs.unsigned_repr; [>| apply Ptrofs.unsigned_range_2]
+
+ | [ |- context[0 <= Ptrofs.unsigned _] ] => apply Ptrofs.unsigned_range_2
+ end.
+
+Module IntExtra.
+ Import Int.
Ltac int_mod_match m :=
match goal with
| [ H : ?x = 0 |- context[?x] ] => rewrite H
@@ -202,7 +216,7 @@ Module IntExtra.
rewrite Zmod_mod
| [ _ : _ |- context[(_ mod Int.modulus) mod m ] ] =>
rewrite <- Zmod_div_mod;
- try (simplify; lia || assumption)
+ try (crush; lia || assumption)
| [ _ : _ |- context[Int.modulus mod m] ] =>
rewrite Zdivide_mod with (a := Int.modulus);
@@ -226,41 +240,130 @@ Module IntExtra.
Ltac int_mod_tac m :=
repeat (int_mod_match m); lia.
- Lemma mul_mod :
+ Lemma mul_mod1 :
+ forall x y m,
+ 0 < m ->
+ (m | Int.modulus) ->
+ Int.unsigned x mod m = 0 ->
+ (Int.unsigned (Int.mul x y)) mod m = 0.
+ Proof.
+ intros. unfold Int.mul.
+ rewrite Int.unsigned_repr_eq.
+
+ repeat match goal with
+ | [ _ : _ |- context[if ?x then _ else _] ] => destruct x
+ | [ _ : _ |- context[_ mod Int.modulus mod m] ] =>
+ rewrite <- Zmod_div_mod; try lia; try assumption
+ end; try (crush; lia); int_mod_tac m.
+ Qed.
+
+ Lemma mul_mod2 :
forall x y m,
0 < m ->
(m | Int.modulus) ->
- Int.signed x mod m = 0 ->
- Int.signed y mod m = 0 ->
- (Int.signed (Int.mul x y)) mod m = 0.
+ Int.unsigned y mod m = 0 ->
+ (Int.unsigned (Int.mul x y)) mod m = 0.
Proof.
intros. unfold Int.mul.
- rewrite Int.signed_repr_eq.
+ rewrite Int.unsigned_repr_eq.
repeat match goal with
| [ _ : _ |- context[if ?x then _ else _] ] => destruct x
| [ _ : _ |- context[_ mod Int.modulus mod m] ] =>
rewrite <- Zmod_div_mod; try lia; try assumption
- | [ _ : _ |- context[Int.unsigned _] ] => rewrite Int.unsigned_signed
- end; try (simplify; lia); int_mod_tac m.
+ end; try (crush; lia); int_mod_tac m.
Qed.
Lemma add_mod :
forall x y m,
0 < m ->
(m | Int.modulus) ->
- Int.signed x mod m = 0 ->
- Int.signed y mod m = 0 ->
- (Int.signed (Int.add x y)) mod m = 0.
+ Int.unsigned x mod m = 0 ->
+ Int.unsigned y mod m = 0 ->
+ (Int.unsigned (Int.add x y)) mod m = 0.
Proof.
intros. unfold Int.add.
- rewrite Int.signed_repr_eq.
+ rewrite Int.unsigned_repr_eq.
repeat match goal with
| [ _ : _ |- context[if ?x then _ else _] ] => destruct x
| [ _ : _ |- context[_ mod Int.modulus mod m] ] =>
rewrite <- Zmod_div_mod; try lia; try assumption
- | [ _ : _ |- context[Int.unsigned _] ] => rewrite Int.unsigned_signed
- end; try (simplify; lia); int_mod_tac m.
+ end; try (crush; lia); int_mod_tac m.
Qed.
+
+ Definition ofbytes (a b c d : byte) : int :=
+ or (shl (repr (Byte.unsigned a)) (repr (3 * Byte.zwordsize)))
+ (or (shl (repr (Byte.unsigned b)) (repr (2 * Byte.zwordsize)))
+ (or (shl (repr (Byte.unsigned c)) (repr Byte.zwordsize))
+ (repr (Byte.unsigned d)))).
+
+ Definition byte0 (n: int) : byte := Byte.repr $ unsigned n.
+ Definition ibyte0 (n: int) : int := Int.repr $ Byte.unsigned $ byte0 n.
+
+ Definition byte1 (n: int) : byte := Byte.repr $ unsigned $ shru n $ repr Byte.zwordsize.
+ Definition ibyte1 (n: int) : int := Int.repr $ Byte.unsigned $ byte1 n.
+
+ Definition byte2 (n: int) : byte := Byte.repr $ unsigned $ shru n $ repr (2 * Byte.zwordsize).
+ Definition ibyte2 (n: int) : int := Int.repr $ Byte.unsigned $ byte2 n.
+
+ Definition byte3 (n: int) : byte := Byte.repr $ unsigned $ shru n $ repr (3 * Byte.zwordsize).
+ Definition ibyte3 (n: int) : int := Int.repr $ Byte.unsigned $ byte3 n.
+
+ Lemma bits_byte0:
+ forall n i, 0 <= i < Byte.zwordsize -> Byte.testbit (byte0 n) i = testbit n i.
+ Proof.
+ intros. unfold byte0. rewrite Byte.testbit_repr; auto.
+ Qed.
+
+ Lemma bits_byte1:
+ forall n i, 0 <= i < Byte.zwordsize -> Byte.testbit (byte1 n) i = testbit n (i + Byte.zwordsize).
+ Proof.
+ intros. unfold byte1. rewrite Byte.testbit_repr; auto.
+ assert (zwordsize = 4 * Byte.zwordsize) by reflexivity.
+ fold (testbit (shru n (repr Byte.zwordsize)) i). rewrite bits_shru.
+ change (unsigned (repr Byte.zwordsize)) with Byte.zwordsize.
+ apply zlt_true. omega. omega.
+ Qed.
+
+ Lemma bits_byte2:
+ forall n i, 0 <= i < Byte.zwordsize -> Byte.testbit (byte2 n) i = testbit n (i + (2 * Byte.zwordsize)).
+ Proof.
+ intros. unfold byte2. rewrite Byte.testbit_repr; auto.
+ assert (zwordsize = 4 * Byte.zwordsize) by reflexivity.
+ fold (testbit (shru n (repr (2 * Byte.zwordsize))) i). rewrite bits_shru.
+ change (unsigned (repr (2 * Byte.zwordsize))) with (2 * Byte.zwordsize).
+ apply zlt_true. omega. omega.
+ Qed.
+
+ Lemma bits_byte3:
+ forall n i, 0 <= i < Byte.zwordsize -> Byte.testbit (byte3 n) i = testbit n (i + (3 * Byte.zwordsize)).
+ Proof.
+ intros. unfold byte3. rewrite Byte.testbit_repr; auto.
+ assert (zwordsize = 4 * Byte.zwordsize) by reflexivity.
+ fold (testbit (shru n (repr (3 * Byte.zwordsize))) i). rewrite bits_shru.
+ change (unsigned (repr (3 * Byte.zwordsize))) with (3 * Byte.zwordsize).
+ apply zlt_true. omega. omega.
+ Qed.
+
+ Lemma bits_ofwords:
+ forall b4 b3 b2 b1 i, 0 <= i < zwordsize ->
+ testbit (ofbytes b4 b3 b2 b1) i =
+ if zlt i Byte.zwordsize
+ then Byte.testbit b1 i
+ else (if zlt i (2 * Byte.zwordsize)
+ then Byte.testbit b2 (i - Byte.zwordsize)
+ else (if zlt i (3 * Byte.zwordsize)
+ then Byte.testbit b2 (i - 2 * Byte.zwordsize)
+ else Byte.testbit b2 (i - 3 * Byte.zwordsize))).
+ Proof.
+ intros. unfold ofbytes. repeat (rewrite bits_or; auto). repeat (rewrite bits_shl; auto).
+ change (unsigned (repr Byte.zwordsize)) with Byte.zwordsize.
+ change (unsigned (repr (2 * Byte.zwordsize))) with (2 * Byte.zwordsize).
+ change (unsigned (repr (3 * Byte.zwordsize))) with (3 * Byte.zwordsize).
+ assert (zwordsize = 4 * Byte.zwordsize) by reflexivity.
+ destruct (zlt i Byte.zwordsize).
+ rewrite testbit_repr; auto.
+ Abort.
+
End IntExtra.
diff --git a/src/common/Monad.v b/src/common/Monad.v
index 8517186..628963e 100644
--- a/src/common/Monad.v
+++ b/src/common/Monad.v
@@ -20,6 +20,10 @@ Module MonadExtra(M : Monad).
Module MonadNotation.
+ Notation "A ; B" :=
+ (bind A (fun _ => B))
+ (at level 200, B at level 200).
+
Notation "'do' X <- A ; B" :=
(bind A (fun X => B))
(at level 200, X ident, A at level 100, B at level 200).
diff --git a/src/common/ZExtra.v b/src/common/ZExtra.v
index a0dd717..519ee7c 100644
--- a/src/common/ZExtra.v
+++ b/src/common/ZExtra.v
@@ -31,4 +31,19 @@ Module ZExtra.
apply Zmult_gt_reg_r in g; lia.
Qed.
+ Lemma Ple_not_eq :
+ forall x y,
+ (x < y)%positive -> x <> y.
+ Proof. lia. Qed.
+
+ Lemma Pge_not_eq :
+ forall x y,
+ (y < x)%positive -> x <> y.
+ Proof. lia. Qed.
+
+ Lemma Ple_Plt_Succ :
+ forall x y n,
+ (x <= y)%positive -> (x < y + n)%positive.
+ Proof. lia. Qed.
+
End ZExtra.
diff --git a/src/extraction/Extraction.v b/src/extraction/Extraction.v
index ba87af6..5d10cd7 100644
--- a/src/extraction/Extraction.v
+++ b/src/extraction/Extraction.v
@@ -16,7 +16,7 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-From coqup Require Verilog Value Compiler.
+From coqup Require Verilog ValueInt Compiler.
From Coq Require DecidableClass.
@@ -128,6 +128,7 @@ Extract Constant Compiler.print_Clight => "PrintClight.print_if".
Extract Constant Compiler.print_Cminor => "PrintCminor.print_if".
Extract Constant driver.Compiler.print_RTL => "PrintRTL.print_if".
Extract Constant Compiler.print_RTL => "PrintRTL.print_if".
+Extract Constant Compiler.print_HTL => "PrintHTL.print_if".
Extract Constant Compiler.print_LTL => "PrintLTL.print_if".
Extract Constant Compiler.print_Mach => "PrintMach.print_if".
Extract Constant Compiler.print => "fun (f: 'a -> unit) (x: 'a) -> f x; x".
@@ -166,7 +167,7 @@ Set Extraction AccessOpaque.
Cd "src/extraction".
Separate Extraction
- Verilog.module Value.uvalueToZ coqup.Compiler.transf_hls
+ Verilog.module ValueInt.uvalueToZ coqup.Compiler.transf_hls
Compiler.transf_c_program Compiler.transf_cminor_program
Cexec.do_initial_state Cexec.do_step Cexec.at_final_state
diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v
index 35b815e..32b6e04 100644
--- a/src/translation/HTLgen.v
+++ b/src/translation/HTLgen.v
@@ -19,7 +19,7 @@
From compcert Require Import Maps.
From compcert Require Errors Globalenvs Integers.
From compcert Require Import AST RTL.
-From coqup Require Import Verilog HTL Coquplib AssocMap Value Statemonad.
+From coqup Require Import Verilog HTL Coquplib AssocMap ValueInt Statemonad.
Hint Resolve AssocMap.gempty : htlh.
Hint Resolve AssocMap.gso : htlh.
@@ -88,10 +88,10 @@ Import HTLMonadExtra.
Export MonadNotation.
Definition state_goto (st : reg) (n : node) : stmnt :=
- Vnonblock (Vvar st) (Vlit (posToValue 32 n)).
+ Vnonblock (Vvar st) (Vlit (posToValue n)).
Definition state_cond (st : reg) (c : expr) (n1 n2 : node) : stmnt :=
- Vnonblock (Vvar st) (Vternary c (posToExpr 32 n1) (posToExpr 32 n2)).
+ Vnonblock (Vvar st) (Vternary c (posToExpr n1) (posToExpr n2)).
Definition check_empty_node_datapath:
forall (s: state) (n: node), { s.(st_datapath)!n = None } + { True }.
@@ -244,7 +244,7 @@ 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)).
+ Vbinop op (Vvar r) (Vlit (ZToValue l)).
Definition translate_comparison (c : Integers.comparison) (args : list reg) : mon expr :=
match c, args with
@@ -294,14 +294,14 @@ Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon ex
| Op.Aindexed off, r1::nil =>
ret (boplitz Vadd r1 off)
| Op.Ascaled scale offset, r1::nil =>
- ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue 32 offset)))
+ ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue offset)))
| Op.Aindexed2 offset, r1::r2::nil =>
ret (Vbinop Vadd (Vvar r1) (boplitz Vadd r2 offset))
| Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *)
ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale))
| Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *)
let a := Integers.Ptrofs.unsigned a in
- ret (Vlit (ZToValue 32 a))
+ ret (Vlit (ZToValue a))
| _, _ => error (Errors.msg "Veriloggen: translate_eff_addressing unsuported addressing")
end.
@@ -332,7 +332,7 @@ Definition translate_instr (op : Op.operation) (args : list reg) : mon expr :=
| Op.Oshr, r1::r2::nil => ret (bop Vshr r1 r2)
| Op.Oshrimm n, r::nil => ret (boplit Vshr r n)
| Op.Oshrximm n, r::nil => ret (Vbinop Vdiv (Vvar r)
- (Vbinop Vshl (Vlit (ZToValue 32 1))
+ (Vbinop Vshl (Vlit (ZToValue 1))
(Vlit (intToValue n))))
| Op.Oshru, r1::r2::nil => ret (bop Vshru r1 r2)
| Op.Oshruimm n, r::nil => ret (boplit Vshru r n)
@@ -343,8 +343,6 @@ Definition translate_instr (op : Op.operation) (args : list reg) : mon expr :=
do tc <- translate_condition c rl;
ret (Vternary tc (Vvar r1) (Vvar r2))
| Op.Olea a, _ => translate_eff_addressing a args
- | Op.Oleal a, _ => translate_eff_addressing a args (* FIXME: Need to be careful here; large arrays might fail? *)
- | Op.Ocast32signed, r::nil => ret (Vvar r) (* FIXME: Don't need to sign extend for now since everything is 32 bit? *)
| _, _ => error (Errors.msg "Htlgen: Instruction not implemented: other")
end.
@@ -382,26 +380,38 @@ Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit :=
| _, _ => Error (Errors.msg "Htlgen: add_branch_instr")
end.
-Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing)
- (args : list reg) (stack : reg) : mon expr :=
- match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*)
- | Mint32, Op.Aindexed off, r1::nil =>
- if (check_address_parameter_signed off)
- then ret (Vvari stack (Vbinop Vdivu (boplitz Vadd r1 off) (Vlit (ZToValue 32 4))))
- else error (Errors.msg "HTLgen: translate_arr_access address misaligned")
- | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *)
- if (check_address_parameter_signed scale) && (check_address_parameter_signed offset)
- then ret (Vvari stack
- (Vbinop Vdivu
- (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale))
- (ZToValue 32 4)))
- else error (Errors.msg "HTLgen: translate_arr_access address misaligned")
- | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *)
+(* Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) *)
+(* (args : list reg) (stack : reg) : mon expr := *)
+(* match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) *)
+(* | Mint32, Op.Aindexed off, r1::nil => *)
+(* if (check_address_parameter_signed off) *)
+(* then ret (Vvari stack (Vbinop Vdivu (boplitz Vadd r1 off) (Vlit (ZToValue 4)))) *)
+(* else error (Errors.msg "HTLgen: translate_arr_access address misaligned") *)
+(* | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) *)
+(* if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) *)
+(* then ret (Vvari stack *)
+(* (Vbinop Vdivu *)
+(* (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) *)
+(* (Vlit (ZToValue 4)))) *)
+(* else error (Errors.msg "HTLgen: translate_arr_access address misaligned") *)
+(* | Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *) *)
+(* let a := Integers.Ptrofs.unsigned a in *)
+(* if (check_address_parameter_unsigned a) *)
+(* then ret (Vvari stack (Vlit (ZToValue (a / 4)))) *)
+(* else error (Errors.msg "HTLgen: eff_addressing misaligned stack offset") *)
+(* | _, _, _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") *)
+(* end. *)
+
+Definition translate_arr_addressing (a: Op.addressing) (args: list reg) : mon expr :=
+ match a, args with (* TODO: We should be more methodical here; what are the possibilities?*)
+ | Op.Aindexed off, r1::nil =>
+ ret (boplitz Vadd r1 off)
+ | Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *)
+ ret (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale))
+ | Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *)
let a := Integers.Ptrofs.unsigned a in
- if (check_address_parameter_unsigned a)
- then ret (Vvari stack (Vlit (ZToValue 32 (a / 4))))
- else error (Errors.msg "HTLgen: eff_addressing misaligned stack offset")
- | _, _, _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing")
+ ret (Vlit (ZToValue a))
+ | _, _ => error (Errors.msg "Veriloggen: translate_arr_addressing unsuported addressing")
end.
Fixpoint enumerate (i : nat) (ns : list node) {struct ns} : list (nat * node) :=
@@ -412,41 +422,70 @@ Fixpoint enumerate (i : nat) (ns : list node) {struct ns} : list (nat * node) :=
Definition tbl_to_case_expr (st : reg) (ns : list node) : list (expr * stmnt) :=
List.map (fun a => match a with
- (i, n) => (Vlit (natToValue 32 i), Vnonblock (Vvar st) (Vlit (posToValue 32 n)))
+ (i, n) => (Vlit (natToValue i), Vnonblock (Vvar st) (Vlit (posToValue n)))
end)
(enumerate 0 ns).
+Definition create_single_cycle_load (stack : reg) (addr : expr) (dst : reg) : stmnt :=
+ Vnonblock (Vvar dst) (Vload stack addr).
+
+Definition create_single_cycle_store (stack : reg) (addr : expr) (src : reg) : stmnt :=
+ let l0 := Vnonblock (Vvari stack addr) (Vvarb0 src) in
+ let l1 := Vnonblock (Vvari stack $ Vbinop Vadd addr (Vlit $ ZToValue 1)) (Vvarb1 src) in
+ let l2 := Vnonblock (Vvari stack $ Vbinop Vadd addr (Vlit $ ZToValue 2)) (Vvarb2 src) in
+ let l3 := Vnonblock (Vvari stack $ Vbinop Vadd addr (Vlit $ ZToValue 3)) (Vvarb3 src)
+ in Vseq l0 $ Vseq l1 $ Vseq l2 $ l3.
+
Definition transf_instr (fin rtrn stack: reg) (ni: node * instruction) : mon unit :=
match ni with
(n, i) =>
match i with
- | Inop n' => add_instr n n' Vskip
+ | Inop n' =>
+ if Z.leb (Z.pos n') Integers.Int.max_unsigned then
+ add_instr n n' Vskip
+ else error (Errors.msg "State is larger than 2^32.")
| Iop op args dst n' =>
- do instr <- translate_instr op args;
- do _ <- declare_reg None dst 32;
- add_instr n n' (nonblock dst instr)
- | Iload mem addr args dst n' =>
- do src <- translate_arr_access mem addr args stack;
- do _ <- declare_reg None dst 32;
- add_instr n n' (nonblock dst src)
- | Istore mem addr args src n' =>
- do dst <- translate_arr_access mem addr args stack;
- add_instr n n' (Vnonblock dst (Vvar src)) (* TODO: Could juse use add_instr? reg exists. *)
+ if Z.leb (Z.pos n') Integers.Int.max_unsigned then
+ do instr <- translate_instr op args;
+ do _ <- declare_reg None dst 32;
+ add_instr n n' (nonblock dst instr)
+ else error (Errors.msg "State is larger than 2^32.")
+ | Iload chunk addr args dst n' =>
+ match chunk with
+ | Mint32 =>
+ if Z.leb (Z.pos n') Integers.Int.max_unsigned
+ then do addr' <- translate_arr_addressing addr args;
+ do _ <- declare_reg None dst 32;
+ add_instr n n' $ create_single_cycle_load stack addr' dst
+ else error (Errors.msg "State is larger than 2^32.")
+ | _ => error (Errors.msg "Iload invalid chunk size.")
+ end
+ | Istore chunk addr args src n' =>
+ match chunk with
+ | Mint32 =>
+ if Z.leb (Z.pos n') Integers.Int.max_unsigned
+ then do addr' <- translate_arr_addressing addr args;
+ add_instr n n' $ create_single_cycle_store stack addr' src
+ else error (Errors.msg "State is larger than 2^32.")
+ | _ => error (Errors.msg "Istore invalid chunk size.")
+ end
| 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
+ if Z.leb (Z.pos n1) Integers.Int.max_unsigned && Z.leb (Z.pos n2) Integers.Int.max_unsigned then
+ do e <- translate_condition cond args;
+ add_branch_instr e n n1 n2
+ else error (Errors.msg "State is larger than 2^32.")
| Ijumptable r tbl =>
do s <- get;
add_node_skip n (Vcase (Vvar r) (tbl_to_case_expr s.(st_st) tbl) (Some Vskip))
| Ireturn r =>
match r with
| Some r' =>
- add_instr_skip n (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r')))
+ add_instr_skip n (Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vvar r')))
| None =>
- add_instr_skip n (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vlit (ZToValue 1%nat 0%Z))))
+ add_instr_skip n (Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vlit (ZToValue 0%Z))))
end
end
end.
@@ -502,32 +541,67 @@ Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) :=
Definition stack_correct (sz : Z) : bool :=
(0 <=? sz) && (sz <? Integers.Ptrofs.modulus) && (Z.modulo sz 4 =? 0).
+Definition max_pc_map (m : Maps.PTree.t stmnt) :=
+ PTree.fold (fun m pc i => Pos.max m pc) m 1%positive.
+
+Lemma max_pc_map_sound:
+ forall m pc i, m!pc = Some i -> Ple pc (max_pc_map m).
+Proof.
+ intros until i. unfold max_pc_function.
+ apply PTree_Properties.fold_rec with (P := fun c m => c!pc = Some i -> Ple pc m).
+ (* extensionality *)
+ intros. apply H0. rewrite H; auto.
+ (* base case *)
+ rewrite PTree.gempty. congruence.
+ (* inductive case *)
+ intros. rewrite PTree.gsspec in H2. destruct (peq pc k).
+ inv H2. xomega.
+ apply Ple_trans with a. auto. xomega.
+Qed.
+
+Lemma max_pc_wf :
+ forall m, Z.pos (max_pc_map m) <= Integers.Int.max_unsigned ->
+ map_well_formed m.
+Proof.
+ unfold map_well_formed. intros.
+ exploit list_in_map_inv. eassumption. intros [x [A B]]. destruct x.
+ apply Maps.PTree.elements_complete in B. apply max_pc_map_sound in B.
+ unfold Ple in B. apply Pos2Z.pos_le_pos in B. subst.
+ simplify. transitivity (Z.pos (max_pc_map m)); eauto.
+Qed.
+
Definition transf_module (f: function) : mon module :=
if stack_correct f.(fn_stacksize) then
- do fin <- create_reg (Some Voutput) 1;
- do rtrn <- create_reg (Some Voutput) 32;
- do (stack, stack_len) <- create_arr None 32 (Z.to_nat (f.(fn_stacksize) / 4));
- do _ <- collectlist (transf_instr fin rtrn stack) (Maps.PTree.elements f.(RTL.fn_code));
- do _ <- collectlist (fun r => declare_reg (Some Vinput) r 32) f.(RTL.fn_params);
- do start <- create_reg (Some Vinput) 1;
- do rst <- create_reg (Some Vinput) 1;
- do clk <- create_reg (Some Vinput) 1;
- do current_state <- get;
- ret (mkmodule
- f.(RTL.fn_params)
- current_state.(st_datapath)
- current_state.(st_controllogic)
- f.(fn_entrypoint)
- current_state.(st_st)
- stack
- stack_len
- fin
- rtrn
- start
- rst
- clk
- current_state.(st_scldecls)
- current_state.(st_arrdecls))
+ do fin <- create_reg (Some Voutput) 1;
+ do rtrn <- create_reg (Some Voutput) 32;
+ do (stack, stack_len) <- create_arr None 8 (Z.to_nat f.(fn_stacksize));
+ do _ <- collectlist (transf_instr fin rtrn stack) (Maps.PTree.elements f.(RTL.fn_code));
+ do _ <- collectlist (fun r => declare_reg (Some Vinput) r 32) f.(RTL.fn_params);
+ do start <- create_reg (Some Vinput) 1;
+ do rst <- create_reg (Some Vinput) 1;
+ do clk <- create_reg (Some Vinput) 1;
+ do current_state <- get;
+ match zle (Z.pos (max_pc_map current_state.(st_datapath))) Integers.Int.max_unsigned,
+ zle (Z.pos (max_pc_map current_state.(st_controllogic))) Integers.Int.max_unsigned with
+ | left LEDATA, left LECTRL =>
+ ret (mkmodule
+ f.(RTL.fn_params)
+ current_state.(st_datapath)
+ current_state.(st_controllogic)
+ f.(fn_entrypoint)
+ current_state.(st_st)
+ stack
+ stack_len
+ fin
+ rtrn
+ start
+ rst
+ clk
+ current_state.(st_scldecls)
+ current_state.(st_arrdecls)
+ (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA)))
+ | _, _ => error (Errors.msg "More than 2^32 states.")
+ end
else error (Errors.msg "Stack size misalignment.").
Definition max_state (f: function) : state :=
diff --git a/src/translation/HTLgenproof.v b/src/translation/HTLgenproof.v
index dfacb5e..51c0fa1 100644
--- a/src/translation/HTLgenproof.v
+++ b/src/translation/HTLgenproof.v
@@ -16,9 +16,9 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-From compcert Require RTL Registers AST Integers.
-From compcert Require Import Globalenvs Memory.
-From coqup Require Import Coquplib HTLgenspec HTLgen Value AssocMap Array IntegerExtra ZExtra.
+From compcert Require RTL Registers AST.
+From compcert Require Import Integers Globalenvs Memory Linking.
+From coqup Require Import Coquplib HTLgenspec HTLgen ValueInt AssocMap Array IntegerExtra ZExtra.
From coqup Require HTL Verilog.
Require Import Lia.
@@ -41,7 +41,7 @@ Hint Constructors match_assocmaps : htlproof.
Definition state_st_wf (m : HTL.module) (s : HTL.state) :=
forall st asa asr res,
s = HTL.State res m st asa asr ->
- asa!(m.(HTL.mod_st)) = Some (posToValue 32 st).
+ asa!(m.(HTL.mod_st)) = Some (posToValue st).
Hint Unfold state_st_wf : htlproof.
Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem : mem) :
@@ -53,7 +53,7 @@ Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem
0 <= ptr < Z.of_nat m.(HTL.mod_stk_len) ->
opt_val_value_lessdef (Mem.loadv AST.Mint32 mem
(Values.Val.offset_ptr sp (Integers.Ptrofs.repr (4 * ptr))))
- (Option.default (NToValue 32 0)
+ (Option.default (NToValue 0)
(Option.join (array_get_error (Z.to_nat ptr) stack)))) ->
match_arrs m f sp mem asa.
@@ -83,39 +83,35 @@ Definition stack_bounds (sp : Values.val) (hi : Z) (m : mem) : Prop :=
Mem.loadv AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr )) = None /\
Mem.storev AST.Mint32 m (Values.Val.offset_ptr sp (Integers.Ptrofs.repr ptr )) v = None.
-Inductive match_stacks (mem : mem) : list RTL.stackframe -> list HTL.stackframe -> Prop :=
-| match_stacks_nil :
- match_stacks mem nil nil
-| match_stacks_cons :
- forall cs lr r f sp sp' pc rs m asr asa
- (TF : tr_module f m)
- (ST: match_stacks mem cs lr)
- (MA: match_assocmaps f rs asr)
- (MARR : match_arrs m f sp mem asa)
- (SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0))
- (RSBP: reg_stack_based_pointers sp' rs)
- (ASBP: arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp)
- (BOUNDS : stack_bounds sp (f.(RTL.fn_stacksize)) mem),
- match_stacks mem (RTL.Stackframe r f sp pc rs :: cs)
- (HTL.Stackframe r m pc asr asa :: lr).
+Inductive match_frames : list RTL.stackframe -> list HTL.stackframe -> Prop :=
+| match_frames_nil :
+ match_frames nil nil.
+
+Inductive match_constants : HTL.module -> assocmap -> Prop :=
+ match_constant :
+ forall m asr,
+ asr!(HTL.mod_reset m) = Some (ZToValue 0) ->
+ asr!(HTL.mod_finish m) = Some (ZToValue 0) ->
+ match_constants m asr.
Inductive match_states : RTL.state -> HTL.state -> Prop :=
| match_state : forall asa asr sf f sp sp' rs mem m st res
(MASSOC : match_assocmaps f rs asr)
(TF : tr_module f m)
(WF : state_st_wf m (HTL.State res m st asr asa))
- (MS : match_stacks mem sf res)
+ (MF : match_frames sf res)
(MARR : match_arrs m f sp mem asa)
(SP : sp = Values.Vptr sp' (Integers.Ptrofs.repr 0))
(RSBP : reg_stack_based_pointers sp' rs)
(ASBP : arr_stack_based_pointers sp' mem (f.(RTL.fn_stacksize)) sp)
- (BOUNDS : stack_bounds sp (f.(RTL.fn_stacksize)) mem),
+ (BOUNDS : stack_bounds sp (f.(RTL.fn_stacksize)) mem)
+ (CONST : match_constants m asr),
match_states (RTL.State sf f sp st rs mem)
(HTL.State res m st asr asa)
| match_returnstate :
forall
v v' stack mem res
- (MS : match_stacks mem stack res),
+ (MF : match_frames stack res),
val_value_lessdef v v' ->
match_states (RTL.Returnstate stack v mem) (HTL.Returnstate res v')
| match_initial_call :
@@ -128,11 +124,30 @@ Definition match_prog (p: RTL.program) (tp: HTL.program) :=
Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp /\
main_is_internal p = true.
-Definition match_prog_matches :
- forall p tp,
- match_prog p tp ->
- Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp.
- Proof. intros. unfold match_prog in H. tauto. Qed.
+Instance TransfHTLLink (tr_fun: RTL.program -> Errors.res HTL.program):
+ TransfLink (fun (p1: RTL.program) (p2: HTL.program) => match_prog p1 p2).
+Proof.
+ red. intros. exfalso. destruct (link_linkorder _ _ _ H) as [LO1 LO2].
+ apply link_prog_inv in H.
+
+ unfold match_prog in *.
+ unfold main_is_internal in *. simplify. repeat (unfold_match H4).
+ repeat (unfold_match H3). simplify.
+ subst. rewrite H0 in *. specialize (H (AST.prog_main p2)).
+ exploit H.
+ apply Genv.find_def_symbol. exists b. split.
+ assumption. apply Genv.find_funct_ptr_iff. eassumption.
+ apply Genv.find_def_symbol. exists b0. split.
+ assumption. apply Genv.find_funct_ptr_iff. eassumption.
+ intros. inv H3. inv H5. destruct H6. inv H5.
+Qed.
+
+Definition match_prog' (p: RTL.program) (tp: HTL.program) :=
+ Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp.
+
+Lemma match_prog_matches :
+ forall p tp, match_prog p tp -> match_prog' p tp.
+Proof. unfold match_prog. tauto. Qed.
Lemma transf_program_match:
forall p tp, HTLgen.transl_program p = Errors.OK tp -> match_prog p tp.
@@ -183,11 +198,11 @@ Lemma list_combine_none :
length l = n ->
list_combine Verilog.merge_cell (list_repeat None n) l = l.
Proof.
- induction n; intros; simplify.
+ induction n; intros; crush.
- symmetry. apply length_zero_iff_nil. auto.
- - destruct l; simplify.
+ - destruct l; crush.
rewrite list_repeat_cons.
- simplify. f_equal.
+ crush. f_equal.
eauto.
Qed.
@@ -198,7 +213,7 @@ Lemma combine_none :
Proof.
intros.
unfold combine.
- simplify.
+ crush.
rewrite <- (arr_wf a) in H.
apply list_combine_none.
@@ -211,12 +226,12 @@ Lemma list_combine_lookup_first :
nth_error l1 n = Some None ->
nth_error (list_combine Verilog.merge_cell l1 l2) n = nth_error l2 n.
Proof.
- induction l1; intros; simplify.
+ induction l1; intros; crush.
rewrite nth_error_nil in H0.
discriminate.
- destruct l2 eqn:EQl2. simplify.
+ destruct l2 eqn:EQl2. crush.
simpl in H. invert H.
destruct n; simpl in *.
invert H0. simpl. reflexivity.
@@ -243,9 +258,9 @@ Lemma list_combine_lookup_second :
nth_error l1 n = Some (Some x) ->
nth_error (list_combine Verilog.merge_cell l1 l2) n = Some (Some x).
Proof.
- induction l1; intros; simplify; auto.
+ induction l1; intros; crush; auto.
- destruct l2 eqn:EQl2. simplify.
+ destruct l2 eqn:EQl2. crush.
simpl in H. invert H.
destruct n; simpl in *.
invert H0. simpl. reflexivity.
@@ -266,32 +281,6 @@ Proof.
assumption.
Qed.
-(* 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 st_greater_than_res :
- forall m res : positive,
- m <> res.
-Admitted.
-
-Lemma finish_not_return :
- forall r f : positive,
- r <> f.
-Admitted.
-
-Lemma finish_not_res :
- forall f r : positive,
- f <> r.
-Admitted.
-
-Lemma greater_than_max_func :
- forall f st,
- Plt (RTL.max_reg_function f) st.
-Proof. Admitted.
-
Ltac inv_state :=
match goal with
MSTATE : match_states _ _ |- _ =>
@@ -335,6 +324,17 @@ Proof.
constructor.
Qed.
+Lemma arr_lookup_some:
+ forall (z : Z) (r0 : Registers.reg) (r : Verilog.reg) (asr : assocmap) (asa : Verilog.assocmap_arr)
+ (stack : Array (option value)) (H5 : asa ! r = Some stack) n,
+ exists x, Verilog.arr_assocmap_lookup asa r n = Some x.
+Proof.
+ intros z r0 r asr asa stack H5 n.
+ eexists.
+ unfold Verilog.arr_assocmap_lookup. rewrite H5. reflexivity.
+Qed.
+Hint Resolve arr_lookup_some : htlproof.
+
Section CORRECTNESS.
Variable prog : RTL.program.
@@ -379,23 +379,161 @@ Section CORRECTNESS.
(Genv.senv_transf_partial TRANSL').
Hint Resolve senv_preserved : htlproof.
+ Lemma ptrofs_inj :
+ forall a b,
+ Ptrofs.unsigned a = Ptrofs.unsigned b -> a = b.
+ Proof.
+ intros. rewrite <- Ptrofs.repr_unsigned. symmetry. rewrite <- Ptrofs.repr_unsigned.
+ rewrite H. auto.
+ Qed.
+
+ Lemma op_stack_based :
+ forall F V sp v m args rs op ge pc' res0 pc f e fin rtrn st stk,
+ tr_instr fin rtrn st stk (RTL.Iop op args res0 pc')
+ (Verilog.Vnonblock (Verilog.Vvar res0) e)
+ (state_goto st pc') ->
+ reg_stack_based_pointers sp rs ->
+ (RTL.fn_code f) ! pc = Some (RTL.Iop op args res0 pc') ->
+ @Op.eval_operation F V ge (Values.Vptr sp Ptrofs.zero) op
+ (map (fun r : positive => Registers.Regmap.get r rs) args) m = Some v ->
+ stack_based v sp.
+ Proof.
+ Ltac solve_no_ptr :=
+ match goal with
+ | H: reg_stack_based_pointers ?sp ?rs |- stack_based (Registers.Regmap.get ?r ?rs) _ =>
+ solve [apply H]
+ | H1: reg_stack_based_pointers ?sp ?rs, H2: Registers.Regmap.get _ _ = Values.Vptr ?b ?i
+ |- context[Values.Vptr ?b _] =>
+ let H := fresh "H" in
+ assert (H: stack_based (Values.Vptr b i) sp) by (rewrite <- H2; apply H1); simplify; solve [auto]
+ | |- context[Registers.Regmap.get ?lr ?lrs] =>
+ destruct (Registers.Regmap.get lr lrs) eqn:?; simplify; auto
+ | |- stack_based (?f _) _ => unfold f
+ | |- stack_based (?f _ _) _ => unfold f
+ | |- stack_based (?f _ _ _) _ => unfold f
+ | |- stack_based (?f _ _ _ _) _ => unfold f
+ | H: ?f _ _ = Some _ |- _ =>
+ unfold f in H; repeat (unfold_match H); inv H
+ | H: ?f _ _ _ _ _ _ = Some _ |- _ =>
+ unfold f in H; repeat (unfold_match H); inv H
+ | H: map (fun r : positive => Registers.Regmap.get r _) ?args = _ |- _ =>
+ destruct args; inv H
+ | |- context[if ?c then _ else _] => destruct c; try discriminate
+ | H: match _ with _ => _ end = Some _ |- _ => repeat (unfold_match H)
+ | |- context[match ?g with _ => _ end] => destruct g; try discriminate
+ | |- _ => simplify; solve [auto]
+ end.
+ intros F V sp v m args rs op g pc' res0 pc f e fin rtrn st stk INSTR RSBP SEL EVAL.
+ inv INSTR. unfold translate_instr in H5.
+ unfold_match H5; repeat (unfold_match H5); repeat (simplify; solve_no_ptr).
+ Qed.
+
+ Lemma int_inj :
+ forall x y,
+ Int.unsigned x = Int.unsigned y ->
+ x = y.
+ Proof.
+ intros. rewrite <- Int.repr_unsigned at 1. rewrite <- Int.repr_unsigned.
+ rewrite <- H. trivial.
+ Qed.
+
Lemma eval_correct :
- forall sp op rs args m v v' e asr asa f s s' i,
+ forall s sp op rs m v e asr asa f f' stk s' i pc res0 pc' args res ml st,
+ match_states (RTL.State stk f sp pc rs m) (HTL.State res ml st asr asa) ->
+ (RTL.fn_code f) ! pc = Some (RTL.Iop op args res0 pc') ->
Op.eval_operation ge sp op
-(List.map (fun r : BinNums.positive => Registers.Regmap.get r rs) args) m = Some v ->
+ (List.map (fun r : BinNums.positive => Registers.Regmap.get r rs) args) m = Some v ->
translate_instr op args s = OK e s' i ->
- val_value_lessdef v v' ->
- Verilog.expr_runp f asr asa e v'.
- Admitted.
+ exists v', Verilog.expr_runp f' asr asa e v' /\ val_value_lessdef v v'.
+ Proof.
+ Ltac eval_correct_tac :=
+ match goal with
+ | |- context[valueToPtr] => unfold valueToPtr
+ | |- context[valueToInt] => unfold valueToInt
+ | |- context[bop] => unfold bop
+ | |- context[boplit] => unfold boplit
+ | |- val_value_lessdef Values.Vundef _ => solve [constructor]
+ | H : val_value_lessdef _ _ |- val_value_lessdef (Values.Vint _) _ => constructor; inv H
+ | |- val_value_lessdef (Values.Vint _) _ => constructor; auto
+ | H : context[RTL.max_reg_function ?f]
+ |- context[_ (Registers.Regmap.get ?r ?rs) (Registers.Regmap.get ?r0 ?rs)] =>
+ let HPle1 := fresh "HPle" in
+ let HPle2 := fresh "HPle" in
+ assert (HPle1 : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ assert (HPle2 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ apply H in HPle1; apply H in HPle2; eexists; split;
+ [econstructor; eauto; constructor; trivial | inv HPle1; inv HPle2; try (constructor; auto)]
+ | H : context[RTL.max_reg_function ?f]
+ |- context[_ (Registers.Regmap.get ?r ?rs) _] =>
+ let HPle1 := fresh "HPle" in
+ assert (HPle1 : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ apply H in HPle1; eexists; split;
+ [econstructor; eauto; constructor; trivial | inv HPle1; try (constructor; auto)]
+ | H : _ :: _ = _ :: _ |- _ => inv H
+ | |- context[match ?d with _ => _ end] => destruct d eqn:?; try discriminate
+ | |- Verilog.expr_runp _ _ _ _ _ => econstructor
+ | |- val_value_lessdef (?f _ _) _ => unfold f
+ | |- val_value_lessdef (?f _) _ => unfold f
+ | H : ?f (Registers.Regmap.get _ _) _ = Some _ |- _ =>
+ unfold f in H; repeat (unfold_match H)
+ | H1 : Registers.Regmap.get ?r ?rs = Values.Vint _, H2 : val_value_lessdef (Registers.Regmap.get ?r ?rs) _
+ |- _ => rewrite H1 in H2; inv H2
+ | |- _ => eexists; split; try constructor; solve [eauto]
+ | H : context[RTL.max_reg_function ?f] |- context[_ (Verilog.Vvar ?r) (Verilog.Vvar ?r0)] =>
+ let HPle1 := fresh "H" in
+ let HPle2 := fresh "H" in
+ assert (HPle1 : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ assert (HPle2 : Ple r0 (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ apply H in HPle1; apply H in HPle2; eexists; split; try constructor; eauto
+ | H : context[RTL.max_reg_function ?f] |- context[Verilog.Vvar ?r] =>
+ let HPle := fresh "H" in
+ assert (HPle : Ple r (RTL.max_reg_function f)) by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ apply H in HPle; eexists; split; try constructor; eauto
+ | |- context[if ?c then _ else _] => destruct c eqn:?; try discriminate
+ end.
+ intros s sp op rs m v e asr asa f f' stk s' i pc pc' res0 args res ml st MSTATE INSTR EVAL TR_INSTR.
+ inv MSTATE. inv MASSOC. unfold translate_instr in TR_INSTR; repeat (unfold_match TR_INSTR); inv TR_INSTR;
+ unfold Op.eval_operation in EVAL; repeat (unfold_match EVAL); inv EVAL;
+ repeat (simplify; eval_correct_tac; unfold valueToInt in *).
+ - pose proof Integers.Ptrofs.agree32_sub as H2; unfold Integers.Ptrofs.agree32 in H2.
+ unfold Ptrofs.of_int. simpl.
+ apply ptrofs_inj. assert (Archi.ptr64 = false) by auto. eapply H2 in H3.
+ rewrite Ptrofs.unsigned_repr. apply H3. replace Ptrofs.max_unsigned with Int.max_unsigned; auto.
+ apply Int.unsigned_range_2.
+ auto. rewrite Ptrofs.unsigned_repr. replace Ptrofs.max_unsigned with Int.max_unsigned; auto.
+ apply Int.unsigned_range_2. rewrite Ptrofs.unsigned_repr. auto.
+ replace Ptrofs.max_unsigned with Int.max_unsigned; auto.
+ apply Int.unsigned_range_2.
+ - pose proof Integers.Ptrofs.agree32_sub as AGR; unfold Integers.Ptrofs.agree32 in AGR.
+ assert (ARCH: Archi.ptr64 = false) by auto. eapply AGR in ARCH.
+ apply int_inj. unfold Ptrofs.to_int. rewrite Int.unsigned_repr.
+ apply ARCH. Search Ptrofs.unsigned. pose proof Ptrofs.unsigned_range_2.
+ replace Ptrofs.max_unsigned with Int.max_unsigned; auto.
+ pose proof Ptrofs.agree32_of_int. unfold Ptrofs.agree32 in H2.
+ eapply H2 in ARCH. apply ARCH.
+ pose proof Ptrofs.agree32_of_int. unfold Ptrofs.agree32 in H2.
+ eapply H2 in ARCH. apply ARCH.
+ - admit. (* mulhs *)
+ - admit. (* mulhu *)
+ - rewrite H0 in Heqb. rewrite H1 in Heqb. discriminate.
+ - rewrite Heqb in Heqb0. discriminate.
+ - rewrite H0 in Heqb. rewrite H1 in Heqb. discriminate.
+ - rewrite Heqb in Heqb0. discriminate.
+ - admit.
+ - admit. (* ror *)
+ - admit. (* addressing *)
+ - admit. (* eval_condition *)
+ - admit. (* select *)
+ Admitted.
Lemma eval_cond_correct :
forall cond (args : list Registers.reg) s1 c s' i rs args m b f asr asa,
- 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 asr asa c (boolToValue 32 b).
+ 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 asr asa c (boolToValue b).
Admitted.
(** The proof of semantic preservation for the translation of instructions
@@ -421,1701 +559,1450 @@ Section CORRECTNESS.
exists asr' asa',
HTL.step tge (HTL.State res m st asr asa) Events.E0 (HTL.State res m st asr' asa').
+ Opaque combine.
- 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, Smallstep.plus HTL.step tge R1 t R2 /\ match_states S2 R2.
+ Ltac tac0 :=
+ match goal with
+ | [ |- context[Verilog.merge_arrs _ _] ] => unfold Verilog.merge_arrs
+ | [ |- context[Verilog.merge_arr] ] => unfold Verilog.merge_arr
+ | [ |- context[Verilog.merge_regs _ _] ] => unfold Verilog.merge_regs; crush; unfold_merge
+ | [ |- context[reg_stack_based_pointers] ] => unfold reg_stack_based_pointers; intros
+ | [ |- context[Verilog.arr_assocmap_set _ _ _ _] ] => unfold Verilog.arr_assocmap_set
+
+ | [ |- context[HTL.empty_stack] ] => unfold HTL.empty_stack
+
+ | [ |- context[_ # ?d <- _ ! ?d] ] => rewrite AssocMap.gss
+ | [ |- context[_ # ?d <- _ ! ?s] ] => rewrite AssocMap.gso
+ | [ |- context[(AssocMap.empty _) ! _] ] => rewrite AssocMap.gempty
+
+ | [ |- context[array_get_error _ (combine Verilog.merge_cell (arr_repeat None _) _)] ] =>
+ rewrite combine_lookup_first
+
+ | [ |- state_st_wf _ _ ] => unfold state_st_wf; inversion 1
+ | [ |- context[match_states _ _] ] => econstructor; auto
+ | [ |- match_arrs _ _ _ _ _ ] => econstructor; auto
+ | [ |- match_assocmaps _ _ _ # _ <- (posToValue _) ] =>
+ apply regs_lessdef_add_greater; [> unfold Plt; lia | assumption]
+
+ | [ H : ?asa ! ?r = Some _ |- Verilog.arr_assocmap_lookup ?asa ?r _ = Some _ ] =>
+ unfold Verilog.arr_assocmap_lookup; setoid_rewrite H; f_equal
+ | [ |- context[(AssocMap.combine _ _ _) ! _] ] =>
+ try (rewrite AssocMap.gcombine; [> | reflexivity])
+
+ | [ |- context[Registers.Regmap.get ?d (Registers.Regmap.set ?d _ _)] ] =>
+ rewrite Registers.Regmap.gss
+ | [ |- context[Registers.Regmap.get ?s (Registers.Regmap.set ?d _ _)] ] =>
+ destruct (Pos.eq_dec s d) as [EQ|EQ];
+ [> rewrite EQ | rewrite Registers.Regmap.gso; auto]
+
+ | [ H : opt_val_value_lessdef _ _ |- _ ] => invert H
+ | [ H : context[Z.of_nat (Z.to_nat _)] |- _ ] => rewrite Z2Nat.id in H; [> solve crush |]
+ | [ H : _ ! _ = Some _ |- _] => setoid_rewrite H
+ end.
+
+ Ltac small_tac := repeat (crush; try array; try ptrofs); crush; auto.
+ Ltac big_tac := repeat (crush; try array; try ptrofs; try tac0); crush; auto.
+
+ Lemma transl_inop_correct:
+ forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive)
+ (rs : RTL.regset) (m : mem) (pc' : RTL.node),
+ (RTL.fn_code f) ! pc = Some (RTL.Inop pc') ->
+ forall R1 : HTL.state,
+ match_states (RTL.State s f sp pc rs m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m) R2.
Proof.
-(* induction 1; intros R1 MSTATE; try inv_state.
- - (* Inop *)
- unfold match_prog in TRANSL.
- econstructor.
- split.
- apply Smallstep.plus_one.
+ intros s f sp pc rs m pc' H R1 MSTATE.
+ inv_state.
+
+ unfold match_prog in TRANSL.
+ econstructor.
+ split.
+ apply Smallstep.plus_one.
+ eapply HTL.step_module; eauto.
+ inv CONST; assumption.
+ inv CONST; assumption.
+ (* processing of state *)
+ econstructor.
+ crush.
+ econstructor.
+ econstructor.
+ econstructor.
+
+ all: invert MARR; big_tac.
+
+ inv CONST; constructor; simplify; rewrite AssocMap.gso; auto; lia.
+
+ Unshelve. auto.
+ Qed.
+ Hint Resolve transl_inop_correct : htlproof.
+
+ Lemma transl_iop_correct:
+ forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive)
+ (rs : Registers.Regmap.t Values.val) (m : mem) (op : Op.operation) (args : list Registers.reg)
+ (res0 : Registers.reg) (pc' : RTL.node) (v : Values.val),
+ (RTL.fn_code f) ! pc = Some (RTL.Iop op args res0 pc') ->
+ Op.eval_operation ge sp op (map (fun r : positive => Registers.Regmap.get r rs) args) m = Some v ->
+ forall R1 : HTL.state,
+ match_states (RTL.State s f sp pc rs m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\
+ match_states (RTL.State s f sp pc' (Registers.Regmap.set res0 v rs) m) R2.
+ Proof.
+ intros s f sp pc rs m op args res0 pc' v H H0 R1 MSTATE.
+ inv_state. inv MARR.
+ exploit eval_correct; eauto. intros. inversion H1. inversion H2.
+ econstructor. split.
+ apply Smallstep.plus_one.
+ eapply HTL.step_module; eauto.
+ inv CONST. assumption.
+ inv CONST. assumption.
+ econstructor; simpl; trivial.
+ constructor; trivial.
+ econstructor; simpl; eauto.
+ simpl. econstructor. econstructor.
+ apply H5. simplify.
+
+ all: big_tac.
+
+ assert (HPle: Ple res0 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_def; eauto; simpl; auto).
+
+ unfold Ple in HPle. lia.
+ apply regs_lessdef_add_match. assumption.
+ apply regs_lessdef_add_greater. unfold Plt; lia. assumption.
+ assert (HPle: Ple res0 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_def; eauto; simpl; auto).
+ unfold Ple in HPle; lia.
+ eapply op_stack_based; eauto.
+ inv CONST. constructor; simplify. rewrite AssocMap.gso. rewrite AssocMap.gso.
+ assumption. lia.
+ assert (HPle: Ple res0 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_def; eauto; simpl; auto).
+ unfold Ple in HPle. lia.
+ rewrite AssocMap.gso. rewrite AssocMap.gso.
+ assumption. lia.
+ assert (HPle: Ple res0 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_def; eauto; simpl; auto).
+ unfold Ple in HPle. lia.
+ Unshelve. trivial.
+ Qed.
+ Hint Resolve transl_iop_correct : htlproof.
+
+ Ltac tac :=
+ repeat match goal with
+ | [ _ : error _ _ = OK _ _ _ |- _ ] => discriminate
+ | [ _ : context[if (?x && ?y) then _ else _] |- _ ] =>
+ let EQ1 := fresh "EQ" in
+ let EQ2 := fresh "EQ" in
+ destruct x eqn:EQ1; destruct y eqn:EQ2; simpl in *
+ | [ _ : context[if ?x then _ else _] |- _ ] =>
+ let EQ := fresh "EQ" in
+ destruct x eqn:EQ; simpl in *
+ | [ H : ret _ _ = _ |- _ ] => invert H
+ | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x
+ end.
+
+ Ltac inv_arr_access :=
+ match goal with
+ | [ _ : translate_arr_access ?chunk ?addr ?args _ _ = OK ?c _ _ |- _] =>
+ destruct c, chunk, addr, args; crush; tac; crush
+ end.
+
+ Lemma transl_iload_correct:
+ forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive)
+ (rs : Registers.Regmap.t Values.val) (m : mem) (chunk : AST.memory_chunk)
+ (addr : Op.addressing) (args : list Registers.reg) (dst : Registers.reg)
+ (pc' : RTL.node) (a v : Values.val),
+ (RTL.fn_code f) ! pc = Some (RTL.Iload chunk addr args dst pc') ->
+ Op.eval_addressing ge sp addr (map (fun r : positive => Registers.Regmap.get r rs) args) = Some a ->
+ Mem.loadv chunk m a = Some v ->
+ forall R1 : HTL.state,
+ match_states (RTL.State s f sp pc rs m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\
+ match_states (RTL.State s f sp pc' (Registers.Regmap.set dst v rs) m) R2.
+ Proof.
+ intros s f sp pc rs m chunk addr args dst pc' a v H H0 H1 R1 MSTATE.
+ inv_state. inv_arr_access.
+
+ + (** Preamble *)
+ invert MARR. crush.
+
+ unfold Op.eval_addressing in H0.
+ destruct (Archi.ptr64) eqn:ARCHI; crush.
+
+ unfold reg_stack_based_pointers in RSBP.
+ pose proof (RSBP r0) as RSBPr0.
+
+ destruct (Registers.Regmap.get r0 rs) eqn:EQr0; crush.
+
+ rewrite ARCHI in H1. crush.
+ subst.
+
+ pose proof MASSOC as MASSOC'.
+ invert MASSOC'.
+ pose proof (H0 r0).
+ assert (HPler0 : Ple r0 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; crush; eauto).
+ apply H6 in HPler0.
+ invert HPler0; try congruence.
+ rewrite EQr0 in H8.
+ invert H8.
+ clear H0. clear H6.
+
+ unfold check_address_parameter_signed in *;
+ unfold check_address_parameter_unsigned in *; crush.
+
+ remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ asr # r0))
+ (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET.
+
+ (** Modular preservation proof *)
+ (*assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE.
+ { rewrite HeqOFFSET.
+ apply PtrofsExtra.add_mod; crush.
+ rewrite Integers.Ptrofs.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush.
+ apply PtrofsExtra.of_int_mod.
+ rewrite Integers.Int.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush. }
+
+ (** Read bounds proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH.
+ { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; crush; auto.
+ unfold stack_bounds in BOUNDS.
+ exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET)); auto.
+ split; try lia; apply Integers.Ptrofs.unsigned_range_2.
+ small_tac. }
+
+ (** Normalisation proof *)
+ assert (Integers.Ptrofs.repr
+ (4 * Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET)
+ as NORMALISE.
+ { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity.
+ rewrite <- PtrofsExtra.mul_unsigned.
+ apply PtrofsExtra.mul_divu; crush; auto. }
+
+ (** Normalised bounds proof *)
+ assert (0 <=
+ Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))
+ < (RTL.fn_stacksize f / 4))
+ as NORMALISE_BOUND.
+ { split.
+ apply Integers.Ptrofs.unsigned_range_2.
+ assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity.
+ unfold Integers.Ptrofs.divu at 2 in H0.
+ rewrite H0. clear H0.
+ rewrite Integers.Ptrofs.unsigned_repr; crush.
+ apply Zmult_lt_reg_r with (p := 4); try lia.
+ repeat rewrite ZLib.div_mul_undo; try lia.
+ apply Z.div_pos; small_tac.
+ apply Z.div_le_upper_bound; small_tac. }
+
+ inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ];
+ clear NORMALISE_BOUND.
+
+ (** Start of proof proper *)
+ eexists. split.
+ eapply Smallstep.plus_one.
eapply HTL.step_module; eauto.
apply assumption_32bit.
- (* processing of state *)
- econstructor.
- simplify.
- econstructor.
- econstructor.
- econstructor.
- simplify.
-
- unfold Verilog.merge_regs.
- unfold_merge. apply AssocMap.gss.
-
- (* prove match_state *)
- rewrite assumption_32bit.
- econstructor; simplify; eauto.
-
- unfold Verilog.merge_regs.
- unfold_merge. simpl. apply regs_lessdef_add_greater. apply greater_than_max_func.
- assumption.
- unfold Verilog.merge_regs.
- unfold state_st_wf. inversion 1. subst. unfold_merge. apply AssocMap.gss.
-
- (* prove match_arrs *)
- invert MARR. simplify.
- unfold HTL.empty_stack. simplify. unfold Verilog.merge_arrs.
- econstructor.
- simplify. repeat split.
-
- rewrite AssocMap.gcombine.
- 2: { reflexivity. }
- rewrite AssocMap.gss.
- unfold Verilog.merge_arr.
- setoid_rewrite H5.
- reflexivity.
-
- rewrite combine_length.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- reflexivity.
-
- unfold arr_repeat. simplify.
- rewrite list_repeat_len; auto.
- intros.
- erewrite array_get_error_equal.
- eauto. apply combine_none.
-
- assumption.
-
- - (* Iop *)
- (* destruct v eqn:?; *)
- (* try ( *)
- (* destruct op eqn:?; inversion H21; simpl in H0; repeat (unfold_match H0); *)
- (* inversion H0; subst; simpl in *; try (unfold_func H4); try (unfold_func H5); *)
- (* try (unfold_func H6); *)
- (* try (unfold Op.eval_addressing32 in H6; repeat (unfold_match H6); inversion H6; *)
- (* unfold_func H3); *)
-
- (* inversion Heql; inversion MASSOC; subst; *)
- (* assert (HPle : Ple r (RTL.max_reg_function f)) *)
- (* by (eapply RTL.max_reg_function_use; eauto; simpl; auto); *)
- (* apply H1 in HPle; inversion HPle; *)
- (* rewrite H2 in *; discriminate *)
- (* ). *)
-
- (* + econstructor. split. *)
- (* apply Smallstep.plus_one. *)
- (* eapply HTL.step_module; eauto. *)
- (* econstructor; simpl; trivial. *)
- (* constructor; trivial. *)
- (* econstructor; simpl; eauto. *)
- (* eapply eval_correct; eauto. constructor. *)
- (* unfold_merge. simpl. *)
- (* rewrite AssocMap.gso. *)
- (* apply AssocMap.gss. *)
- (* apply st_greater_than_res. *)
-
- (* (* match_states *) *)
- (* assert (pc' = valueToPos (posToValue 32 pc')). auto using assumption_32bit. *)
- (* rewrite <- H1. *)
- (* constructor; auto. *)
- (* unfold_merge. *)
- (* apply regs_lessdef_add_match. *)
- (* constructor. *)
- (* apply regs_lessdef_add_greater. *)
- (* apply greater_than_max_func. *)
- (* assumption. *)
-
- (* unfold state_st_wf. intros. inversion H2. subst. *)
- (* unfold_merge. *)
- (* rewrite AssocMap.gso. *)
- (* apply AssocMap.gss. *)
- (* apply st_greater_than_res. *)
-
- (* + econstructor. split. *)
- (* apply Smallstep.plus_one. *)
- (* eapply HTL.step_module; eauto. *)
- (* econstructor; simpl; trivial. *)
- (* constructor; trivial. *)
- (* econstructor; simpl; eauto. *)
- (* eapply eval_correct; eauto. *)
- (* constructor. rewrite valueToInt_intToValue. trivial. *)
- (* unfold_merge. simpl. *)
- (* rewrite AssocMap.gso. *)
- (* apply AssocMap.gss. *)
- (* apply st_greater_than_res. *)
-
- (* (* match_states *) *)
- (* assert (pc' = valueToPos (posToValue 32 pc')). auto using assumption_32bit. *)
- (* rewrite <- H1. *)
- (* constructor. *)
- (* unfold_merge. *)
- (* apply regs_lessdef_add_match. *)
- (* constructor. *)
- (* symmetry. apply valueToInt_intToValue. *)
- (* apply regs_lessdef_add_greater. *)
- (* apply greater_than_max_func. *)
- (* assumption. assumption. *)
-
- (* unfold state_st_wf. intros. inversion H2. subst. *)
- (* unfold_merge. *)
- (* rewrite AssocMap.gso. *)
- (* apply AssocMap.gss. *)
- (* apply st_greater_than_res. *)
- (* assumption. *)
- admit.
-
- Ltac rt :=
- repeat match goal with
- | [ _ : error _ _ = OK _ _ _ |- _ ] => discriminate
- | [ _ : context[if (?x && ?y) then _ else _] |- _ ] =>
- let EQ1 := fresh "EQ" in
- let EQ2 := fresh "EQ" in
- destruct x eqn:EQ1; destruct y eqn:EQ2; simpl in *
- | [ _ : context[if ?x then _ else _] |- _ ] =>
- let EQ := fresh "EQ" in
- destruct x eqn:EQ; simpl in *
- | [ H : ret _ _ = _ |- _ ] => invert H
- | [ _ : context[match ?x with | _ => _ end] |- _ ] => destruct x
- end.
-
- - (* FIXME: Should be able to use the spec to avoid destructing here? *)
- destruct c, chunk, addr, args; simplify; rt; simplify.
-
- + (** Preamble *)
- invert MARR. simplify.
-
- unfold Op.eval_addressing in H0.
- destruct (Archi.ptr64) eqn:ARCHI; simplify.
-
- unfold reg_stack_based_pointers in RSBP.
- pose proof (RSBP r0) as RSBPr0.
-
- destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify.
-
- rewrite ARCHI in H1. simplify.
- subst.
-
- pose proof MASSOC as MASSOC'.
- invert MASSOC'.
- pose proof (H0 r0).
- assert (HPler0 : Ple r0 (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simplify; eauto).
- apply H6 in HPler0.
- invert HPler0; try congruence.
- rewrite EQr0 in H8.
- invert H8.
- clear H0. clear H6.
-
- unfold check_address_parameter_signed in *;
- unfold check_address_parameter_unsigned in *; simplify.
-
- remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0))
- (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET.
-
- (** Modular preservation proof *)
- assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE.
- { rewrite HeqOFFSET.
- apply PtrofsExtra.add_mod; simplify; try lia.
- exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *)
- rewrite Integers.Ptrofs.signed_repr; try assumption.
- admit. (* FIXME: Register bounds. *)
- apply PtrofsExtra.of_int_mod.
- rewrite Integers.Int.signed_repr; simplify; try split; try assumption.
- }
-
- (** Read bounds proof *)
- assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH.
- { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; simplify; auto.
- unfold stack_bounds in BOUNDS.
- exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET)); auto.
- split; try lia; apply Integers.Ptrofs.unsigned_range_2.
- simplify.
- replace (Integers.Ptrofs.repr 0) with (Integers.Ptrofs.zero) by reflexivity.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr. intros. simplify. congruence.
- apply Integers.Ptrofs.unsigned_range_2. }
-
- (** Normalisation proof *)
- assert (Integers.Ptrofs.repr
- (4 * Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET)
- as NORMALISE.
- { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity.
- rewrite <- PtrofsExtra.mul_unsigned.
- apply PtrofsExtra.mul_divu; simplify; auto; lia. }
-
- (** Normalised bounds proof *)
- assert (0 <=
- Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))
- < (RTL.fn_stacksize f / 4))
- as NORMALISE_BOUND.
- { split.
- apply Integers.Ptrofs.unsigned_range_2.
- assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity.
- unfold Integers.Ptrofs.divu at 2 in H0.
- rewrite H0. clear H0.
- rewrite Integers.Ptrofs.unsigned_repr; simplify.
- apply Zmult_lt_reg_r with (p := 4); try lia.
- repeat rewrite ZLib.div_mul_undo; try lia.
- split.
- apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2.
- apply Z.div_le_upper_bound; lia. }
-
- inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ];
- clear NORMALISE_BOUND.
-
- eexists. split.
- eapply Smallstep.plus_one.
- eapply HTL.step_module; eauto.
- apply assumption_32bit.
- econstructor. econstructor. econstructor. simplify.
- econstructor. econstructor. econstructor. simplify.
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ1]). (* FIXME: These will be shelved and cause sadness. *)
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ2]).
- econstructor. econstructor. econstructor. econstructor.
- econstructor. econstructor. econstructor. econstructor.
-
- all: simplify.
-
- (** Verilog array lookup *)
- unfold Verilog.arr_assocmap_lookup. setoid_rewrite H5.
- f_equal.
-
- (** State Lookup *)
- unfold Verilog.merge_regs.
- simplify.
- unfold_merge.
- rewrite AssocMap.gso.
- apply AssocMap.gss.
- apply st_greater_than_res.
-
- (** Match states *)
- rewrite assumption_32bit.
- econstructor; eauto.
-
- (** Match assocmaps *)
- unfold Verilog.merge_regs. simplify. unfold_merge.
- apply regs_lessdef_add_match.
-
- (** Equality proof *)
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- rewrite ZERO in H7. clear ZERO.
- setoid_rewrite Integers.Ptrofs.add_zero_l in H7.
-
- specialize (H7 (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu
- OFFSET
- (Integers.Ptrofs.repr 4)))).
-
- exploit H7.
- rewrite Z2Nat.id; eauto.
- apply Z.div_pos; lia.
-
- intros I.
-
+ econstructor. econstructor. econstructor. crush.
+ econstructor. econstructor. econstructor. crush.
+ econstructor. econstructor.
+ econstructor. econstructor. econstructor. econstructor.
+ econstructor. econstructor. econstructor. econstructor.
+
+ all: big_tac.
+
+ 1: {
+ assert (HPle : Ple dst (RTL.max_reg_function f)).
+ eapply RTL.max_reg_function_def. eassumption. auto.
+ apply ZExtra.Pge_not_eq. apply ZExtra.Ple_Plt_Succ. assumption.
+ }
+
+ 2: {
+ assert (HPle : Ple dst (RTL.max_reg_function f)).
+ eapply RTL.max_reg_function_def. eassumption. auto.
+ apply ZExtra.Pge_not_eq. apply ZExtra.Ple_Plt_Succ. assumption.
+ }
+
+ (** Match assocmaps *)
+ apply regs_lessdef_add_match; big_tac.
+
+ (** Equality proof *)
+ match goal with
+ | [ |- context [valueToNat ?x] ] =>
assert (Z.to_nat
(Integers.Ptrofs.unsigned
(Integers.Ptrofs.divu
OFFSET
(Integers.Ptrofs.repr 4)))
=
- valueToNat (vdiv (vplus asr # r0 (ZToValue 32 z) ?EQ2) (ZToValue 32 4) ?EQ1))
- as EXPR_OK by admit.
- rewrite <- EXPR_OK.
- rewrite NORMALISE in I.
- rewrite H1 in I.
- invert I. assumption.
-
- (** PC match *)
- apply regs_lessdef_add_greater.
- apply greater_than_max_func.
- assumption.
-
- (** States well formed *)
- unfold state_st_wf. inversion 1. simplify.
- unfold Verilog.merge_regs.
- unfold_merge. rewrite AssocMap.gso.
- apply AssocMap.gss.
- apply st_greater_than_res.
-
- (** Match arrays *)
- econstructor.
- repeat split; simplify.
- unfold HTL.empty_stack.
- simplify.
- unfold Verilog.merge_arrs.
-
- rewrite AssocMap.gcombine.
- 2: { reflexivity. }
- rewrite AssocMap.gss.
- unfold Verilog.merge_arr.
- setoid_rewrite H5.
- reflexivity.
-
- rewrite combine_length.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- reflexivity.
-
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- congruence.
-
- intros.
- erewrite array_get_error_equal.
- eauto. apply combine_none.
- assumption.
-
- (** RSBP preservation *)
- unfold reg_stack_based_pointers. intros.
- destruct (Pos.eq_dec r1 dst); try rewrite e. (* FIXME: Prepare this for automation *)
-
- rewrite Registers.Regmap.gss.
- unfold arr_stack_based_pointers in ASBP.
- specialize (ASBP (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))).
- exploit ASBP; auto; intros I.
-
- rewrite NORMALISE in I.
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- rewrite ZERO in I. clear ZERO.
- simplify.
- rewrite Integers.Ptrofs.add_zero_l in I.
- rewrite H1 in I.
- assumption.
- simplify.
-
- rewrite Registers.Regmap.gso; auto.
-
- + (** Preamble *)
- invert MARR. simplify.
-
- unfold Op.eval_addressing in H0.
- destruct (Archi.ptr64) eqn:ARCHI; simplify.
-
- unfold reg_stack_based_pointers in RSBP.
- pose proof (RSBP r0) as RSBPr0.
- pose proof (RSBP r1) as RSBPr1.
-
- destruct (Registers.Regmap.get r0 rs) eqn:EQr0;
- destruct (Registers.Regmap.get r1 rs) eqn:EQr1; simplify.
-
- rewrite ARCHI in H1. simplify.
- subst.
- clear RSBPr1.
-
- pose proof MASSOC as MASSOC'.
- invert MASSOC'.
- pose proof (H0 r0).
- pose proof (H0 r1).
- assert (HPler0 : Ple r0 (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simplify; eauto).
- assert (HPler1 : Ple r1 (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simpl; auto).
- apply H6 in HPler0.
- apply H8 in HPler1.
- invert HPler0; invert HPler1; try congruence.
- rewrite EQr0 in H10.
- rewrite EQr1 in H12.
- invert H10. invert H12.
- clear H0. clear H6. clear H8.
-
- unfold check_address_parameter_signed in *;
- unfold check_address_parameter_unsigned in *; simplify.
-
- remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0))
- (Integers.Ptrofs.of_int
- (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z))
- (Integers.Int.repr z0)))) as OFFSET.
-
- (** Modular preservation proof *)
- assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE.
- { rewrite HeqOFFSET.
- apply PtrofsExtra.add_mod; simplify; try lia.
- exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *)
- rewrite Integers.Ptrofs.signed_repr; try assumption.
- admit. (* FIXME: Register bounds. *)
- apply PtrofsExtra.of_int_mod.
- apply IntExtra.add_mod; simplify; try lia.
- exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *)
- apply IntExtra.mul_mod; simplify; try lia.
- exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *)
- admit. (* FIXME: Register bounds. *)
- rewrite Integers.Int.signed_repr; simplify; try split; try assumption.
- rewrite Integers.Int.signed_repr; simplify; try split; try assumption.
- }
-
- (** Read bounds proof *)
- assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH.
- { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; simplify; auto.
- unfold stack_bounds in BOUNDS.
- exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET)); auto.
- split; try lia; apply Integers.Ptrofs.unsigned_range_2.
- simplify.
- replace (Integers.Ptrofs.repr 0) with (Integers.Ptrofs.zero) by reflexivity.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr. intros. simplify. congruence.
- apply Integers.Ptrofs.unsigned_range_2. }
-
- (** Normalisation proof *)
- assert (Integers.Ptrofs.repr
- (4 * Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET)
- as NORMALISE.
- { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity.
- rewrite <- PtrofsExtra.mul_unsigned.
- apply PtrofsExtra.mul_divu; simplify; auto; lia. }
-
- (** Normalised bounds proof *)
- assert (0 <=
- Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))
- < (RTL.fn_stacksize f / 4))
- as NORMALISE_BOUND.
- { split.
- apply Integers.Ptrofs.unsigned_range_2.
- assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity.
- unfold Integers.Ptrofs.divu at 2 in H0.
- rewrite H0. clear H0.
- rewrite Integers.Ptrofs.unsigned_repr; simplify.
- apply Zmult_lt_reg_r with (p := 4); try lia.
- repeat rewrite ZLib.div_mul_undo; try lia.
- split.
- apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2.
- apply Z.div_le_upper_bound; lia. }
-
- inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ];
- clear NORMALISE_BOUND.
-
- (** Start of proof proper *)
- eexists. split.
- eapply Smallstep.plus_one.
- eapply HTL.step_module; eauto.
- apply assumption_32bit.
- econstructor. econstructor. econstructor. simplify.
- econstructor. econstructor. econstructor. simplify.
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ3]). (* FIXME: These will be shelved and cause sadness. *)
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ4]).
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ5]).
- econstructor. econstructor. econstructor. econstructor.
- econstructor.
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ6]).
- econstructor. econstructor. econstructor. econstructor.
- econstructor. econstructor. econstructor. econstructor.
- econstructor. econstructor.
-
- all: simplify.
-
- (** Verilog array lookup *)
- unfold Verilog.arr_assocmap_lookup. setoid_rewrite H5.
- f_equal.
-
- (** State Lookup *)
- unfold Verilog.merge_regs.
- simplify.
- unfold_merge.
- rewrite AssocMap.gso.
- apply AssocMap.gss.
- apply st_greater_than_res.
-
- (** Match states *)
- rewrite assumption_32bit.
- econstructor; eauto.
-
- (** Match assocmaps *)
- unfold Verilog.merge_regs. simplify. unfold_merge.
- apply regs_lessdef_add_match.
-
- (** Equality proof *)
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- rewrite ZERO in H7. clear ZERO.
- setoid_rewrite Integers.Ptrofs.add_zero_l in H7.
-
- specialize (H7 (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu
- OFFSET
- (Integers.Ptrofs.repr 4)))).
-
- exploit H7.
- rewrite Z2Nat.id; eauto.
- apply Z.div_pos; lia.
-
- intros I.
- assert (Z.to_nat
- (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu
- OFFSET
- (Integers.Ptrofs.repr 4)))
- = valueToNat
- (vdiv (vplus (vplus asr # r0 (ZToValue 32 z0) ?EQ5)
- (vmul asr # r1 (ZToValue 32 z) ?EQ6) ?EQ4) (ZToValue 32 4) ?EQ3))
- as EXPR_OK by admit.
- rewrite <- EXPR_OK.
- rewrite NORMALISE in I.
- rewrite H1 in I.
- invert I. assumption.
-
- (** PC match *)
- apply regs_lessdef_add_greater.
- apply greater_than_max_func.
- assumption.
-
- (** States well formed *)
- unfold state_st_wf. inversion 1. simplify.
- unfold Verilog.merge_regs.
- unfold_merge. rewrite AssocMap.gso.
- apply AssocMap.gss.
- apply st_greater_than_res.
-
- (** Match arrays *)
- econstructor.
- repeat split; simplify.
- unfold HTL.empty_stack.
- simplify.
- unfold Verilog.merge_arrs.
-
- rewrite AssocMap.gcombine.
- 2: { reflexivity. }
- rewrite AssocMap.gss.
- unfold Verilog.merge_arr.
- setoid_rewrite H5.
- reflexivity.
-
- rewrite combine_length.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- reflexivity.
-
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- congruence.
-
- intros.
- erewrite array_get_error_equal.
- eauto. apply combine_none.
- assumption.
-
- (** RSBP preservation *)
- unfold reg_stack_based_pointers. intros.
- destruct (Pos.eq_dec r2 dst); try rewrite e. (* FIXME: Prepare this for automation *)
-
- rewrite Registers.Regmap.gss.
- unfold arr_stack_based_pointers in ASBP.
- specialize (ASBP (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))).
- exploit ASBP; auto; intros I.
-
- rewrite NORMALISE in I.
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- rewrite ZERO in I. clear ZERO.
- simplify.
- rewrite Integers.Ptrofs.add_zero_l in I.
- rewrite H1 in I.
- assumption.
- simplify.
-
- rewrite Registers.Regmap.gso; auto.
-
- + invert MARR. simplify.
-
- unfold Op.eval_addressing in H0.
- destruct (Archi.ptr64) eqn:ARCHI; simplify.
- rewrite ARCHI in H0. simplify.
-
- unfold check_address_parameter_unsigned in *;
- unfold check_address_parameter_signed in *; simplify.
-
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- rewrite ZERO in H1. clear ZERO.
- rewrite Integers.Ptrofs.add_zero_l in H1.
-
- remember i0 as OFFSET.
-
- (** Modular preservation proof *)
- rename H0 into MOD_PRESERVE.
-
- (** Read bounds proof *)
- assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH.
- { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; simplify; auto.
- unfold stack_bounds in BOUNDS.
- exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET)); auto.
- simplify.
- replace (Integers.Ptrofs.repr 0) with (Integers.Ptrofs.zero) by reflexivity.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr. intros. simplify. congruence.
- apply Integers.Ptrofs.unsigned_range_2. }
-
- (** Normalisation proof *)
- assert (Integers.Ptrofs.repr
- (4 * Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET)
- as NORMALISE.
- { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity.
- rewrite <- PtrofsExtra.mul_unsigned.
- apply PtrofsExtra.mul_divu; simplify; auto; try lia. }
-
- (** Normalised bounds proof *)
- assert (0 <=
- Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))
- < (RTL.fn_stacksize f / 4))
+ valueToNat x)
+ as EXPR_OK by admit
+ end.
+ rewrite <- EXPR_OK.
+
+ specialize (H7 (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu
+ OFFSET
+ (Integers.Ptrofs.repr 4)))).
+ exploit H7; big_tac.
+
+ (** RSBP preservation *)
+ unfold arr_stack_based_pointers in ASBP.
+ specialize (ASBP (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))).
+ exploit ASBP; big_tac.
+ rewrite NORMALISE in H0. rewrite H1 in H0. assumption.
+
+ + (** Preamble *)
+ invert MARR. crush.
+
+ unfold Op.eval_addressing in H0.
+ destruct (Archi.ptr64) eqn:ARCHI; crush.
+
+ unfold reg_stack_based_pointers in RSBP.
+ pose proof (RSBP r0) as RSBPr0.
+ pose proof (RSBP r1) as RSBPr1.
+
+ destruct (Registers.Regmap.get r0 rs) eqn:EQr0;
+ destruct (Registers.Regmap.get r1 rs) eqn:EQr1; crush.
+
+ rewrite ARCHI in H1. crush.
+ subst.
+ clear RSBPr1.
+
+ pose proof MASSOC as MASSOC'.
+ invert MASSOC'.
+ pose proof (H0 r0).
+ pose proof (H0 r1).
+ assert (HPler0 : Ple r0 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; crush; eauto).
+ assert (HPler1 : Ple r1 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; simpl; auto).
+ apply H6 in HPler0.
+ apply H8 in HPler1.
+ invert HPler0; invert HPler1; try congruence.
+ rewrite EQr0 in H9.
+ rewrite EQr1 in H11.
+ invert H9. invert H11.
+ clear H0. clear H6. clear H8.
+
+ unfold check_address_parameter_signed in *;
+ unfold check_address_parameter_unsigned in *; crush.
+
+ remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ asr # r0))
+ (Integers.Ptrofs.of_int
+ (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z))
+ (Integers.Int.repr z0)))) as OFFSET.
+
+ (** Modular preservation proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE.
+ { rewrite HeqOFFSET.
+ apply PtrofsExtra.add_mod; crush; try lia.
+ rewrite Integers.Ptrofs.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush.
+ apply PtrofsExtra.of_int_mod.
+ apply IntExtra.add_mod; crush.
+ apply IntExtra.mul_mod2; crush.
+ rewrite Integers.Int.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush.
+ rewrite Integers.Int.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush. }
+
+ (** Read bounds proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH.
+ { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; crush; auto.
+ unfold stack_bounds in BOUNDS.
+ exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET)); auto.
+ split; try lia; apply Integers.Ptrofs.unsigned_range_2.
+ small_tac. }
+
+ (** Normalisation proof *)
+ assert (Integers.Ptrofs.repr
+ (4 * Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET)
+ as NORMALISE.
+ { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity.
+ rewrite <- PtrofsExtra.mul_unsigned.
+ apply PtrofsExtra.mul_divu; crush. }
+
+ (** Normalised bounds proof *)
+ assert (0 <=
+ Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))
+ < (RTL.fn_stacksize f / 4))
as NORMALISE_BOUND.
- { split.
- apply Integers.Ptrofs.unsigned_range_2.
- assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity.
- unfold Integers.Ptrofs.divu at 2 in H0.
- rewrite H0. clear H0.
- rewrite Integers.Ptrofs.unsigned_repr; simplify.
- apply Zmult_lt_reg_r with (p := 4); try lia.
- repeat rewrite ZLib.div_mul_undo; try lia.
- split.
- apply Z.div_pos; try lia; apply Integers.Ptrofs.unsigned_range_2.
- apply Z.div_le_upper_bound; lia. }
-
- inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ];
- clear NORMALISE_BOUND.
-
- (** Start of proof proper *)
- eexists. split.
- eapply Smallstep.plus_one.
- eapply HTL.step_module; eauto.
- apply assumption_32bit.
- econstructor. econstructor. econstructor. simplify.
- econstructor. econstructor. econstructor. econstructor. simplify.
-
- all: simplify.
-
- (** Verilog array lookup *)
- unfold Verilog.arr_assocmap_lookup. setoid_rewrite H5.
- f_equal.
-
- (** State Lookup *)
- unfold Verilog.merge_regs.
- simplify.
- unfold_merge.
- rewrite AssocMap.gso.
- apply AssocMap.gss.
- apply st_greater_than_res.
-
- (** Match states *)
- rewrite assumption_32bit.
- econstructor; eauto.
-
- (** Match assocmaps *)
- unfold Verilog.merge_regs. simplify. unfold_merge.
- apply regs_lessdef_add_match.
-
- (** Equality proof *)
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- rewrite ZERO in H7. clear ZERO.
- setoid_rewrite Integers.Ptrofs.add_zero_l in H7.
-
- specialize (H7 (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu
- OFFSET
- (Integers.Ptrofs.repr 4)))).
-
- exploit H7.
- rewrite Z2Nat.id; eauto.
- apply Z.div_pos; lia.
-
- intros I.
- assert (Z.to_nat
- (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu
- OFFSET
- (Integers.Ptrofs.repr 4)))
- =
- valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned OFFSET / 4)))
- as EXPR_OK by admit.
- rewrite <- EXPR_OK.
- rewrite NORMALISE in I.
- rewrite H1 in I.
- invert I. assumption.
-
- (** PC match *)
- apply regs_lessdef_add_greater.
- apply greater_than_max_func.
- assumption.
-
- (** States well formed *)
- unfold state_st_wf. inversion 1. simplify.
- unfold Verilog.merge_regs.
- unfold_merge. rewrite AssocMap.gso.
- apply AssocMap.gss.
- apply st_greater_than_res.
-
- (** Match arrays *)
- econstructor.
- repeat split; simplify.
- unfold HTL.empty_stack.
- simplify.
- unfold Verilog.merge_arrs.
-
- rewrite AssocMap.gcombine.
- 2: { reflexivity. }
- rewrite AssocMap.gss.
- unfold Verilog.merge_arr.
- setoid_rewrite H5.
- reflexivity.
-
- rewrite combine_length.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- reflexivity.
-
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- congruence.
-
- intros.
- erewrite array_get_error_equal.
- eauto. apply combine_none.
- assumption.
-
- (** RSBP preservation *)
- unfold reg_stack_based_pointers. intros.
- destruct (Pos.eq_dec r0 dst); try rewrite e. (* FIXME: Prepare this for automation *)
-
- rewrite Registers.Regmap.gss.
- unfold arr_stack_based_pointers in ASBP.
- specialize (ASBP (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))).
- exploit ASBP; auto; intros I.
-
- rewrite NORMALISE in I.
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- rewrite ZERO in I. clear ZERO.
- simplify.
- rewrite Integers.Ptrofs.add_zero_l in I.
- rewrite H1 in I.
- assumption.
- simplify.
-
- rewrite Registers.Regmap.gso; auto.
-
- - destruct c, chunk, addr, args; simplify; rt; simplify.
- + (** Preamble *)
- invert MARR. simplify.
-
- unfold Op.eval_addressing in H0.
- destruct (Archi.ptr64) eqn:ARCHI; simplify.
-
- unfold reg_stack_based_pointers in RSBP.
- pose proof (RSBP r0) as RSBPr0.
-
- destruct (Registers.Regmap.get r0 rs) eqn:EQr0; simplify.
-
- rewrite ARCHI in H1. simplify.
- subst.
-
- pose proof MASSOC as MASSOC'.
- invert MASSOC'.
- pose proof (H0 r0).
- assert (HPler0 : Ple r0 (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simplify; eauto).
- apply H6 in HPler0.
- invert HPler0; try congruence.
- rewrite EQr0 in H8.
- invert H8.
- clear H0. clear H6.
-
- unfold check_address_parameter_unsigned in *;
- unfold check_address_parameter_signed in *; simplify.
-
- remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0))
- (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET.
-
- (** Modular preservation proof *)
- assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE.
- { rewrite HeqOFFSET.
- apply PtrofsExtra.add_mod; simplify; try lia.
- exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *)
- rewrite Integers.Ptrofs.signed_repr; try assumption.
- admit. (* FIXME: Register bounds. *)
- apply PtrofsExtra.of_int_mod.
- rewrite Integers.Int.signed_repr; simplify; try split; try assumption.
- }
-
- (** Write bounds proof *)
- assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH.
- { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; simplify; auto.
- unfold stack_bounds in BOUNDS.
- exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET) (Registers.Regmap.get src rs)); auto.
- split; try lia; apply Integers.Ptrofs.unsigned_range_2.
- simplify.
- replace (Integers.Ptrofs.repr 0) with (Integers.Ptrofs.zero) by reflexivity.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr. intros. simplify. congruence.
- apply Integers.Ptrofs.unsigned_range_2. }
-
- (** Start of proof proper *)
- eexists. split.
- eapply Smallstep.plus_one.
- eapply HTL.step_module; eauto.
- apply assumption_32bit.
- econstructor. econstructor. econstructor.
- eapply Verilog.stmnt_runp_Vnonblock_arr. simplify.
- econstructor.
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ7]).
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ8]).
- econstructor.
- econstructor.
- econstructor. econstructor. econstructor. econstructor.
- econstructor. econstructor. econstructor. econstructor.
-
- all: simplify.
-
- (** State Lookup *)
- unfold Verilog.merge_regs.
- simplify.
- unfold_merge.
- apply AssocMap.gss.
-
- (** Match states *)
- rewrite assumption_32bit.
- econstructor; eauto.
-
- (** Match assocmaps *)
- unfold Verilog.merge_regs. simplify. unfold_merge.
- apply regs_lessdef_add_greater. apply greater_than_max_func.
- assumption.
-
- (** States well formed *)
- unfold state_st_wf. inversion 1. simplify.
- unfold Verilog.merge_regs.
- unfold_merge.
- apply AssocMap.gss.
-
- (** Match stacks *)
- admit.
-
- (** Equality proof *)
+ { split.
+ apply Integers.Ptrofs.unsigned_range_2.
+ assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity.
+ unfold Integers.Ptrofs.divu at 2 in H0.
+ rewrite H0. clear H0.
+ rewrite Integers.Ptrofs.unsigned_repr; crush.
+ apply Zmult_lt_reg_r with (p := 4); try lia.
+ repeat rewrite ZLib.div_mul_undo; try lia.
+ apply Z.div_pos; small_tac.
+ apply Z.div_le_upper_bound; lia. }
+
+ inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ];
+ clear NORMALISE_BOUND.
+
+ (** Start of proof proper *)
+ eexists. split.
+ eapply Smallstep.plus_one.
+ eapply HTL.step_module; eauto.
+ apply assumption_32bit.
+ econstructor. econstructor. econstructor. crush.
+ econstructor. econstructor. econstructor. crush.
+ econstructor. econstructor. econstructor.
+ econstructor. econstructor. econstructor. econstructor.
+ econstructor.
+ eapply Verilog.erun_Vbinop with (EQ := ?[EQ6]).
+ econstructor. econstructor. econstructor. econstructor.
+ econstructor. econstructor. econstructor. econstructor.
+ econstructor. econstructor.
+
+ all: big_tac.
+
+ 1: { assert (HPle : Ple dst (RTL.max_reg_function f)).
+ eapply RTL.max_reg_function_def. eassumption. auto.
+ apply ZExtra.Pge_not_eq. apply ZExtra.Ple_Plt_Succ. assumption. }
+
+ 2: { assert (HPle : Ple dst (RTL.max_reg_function f)).
+ eapply RTL.max_reg_function_def. eassumption. auto.
+ apply ZExtra.Pge_not_eq. apply ZExtra.Ple_Plt_Succ. assumption. }
+
+ (** Match assocmaps *)
+ apply regs_lessdef_add_match; big_tac.
+
+ (** Equality proof *)
+ match goal with
+ | [ |- context [valueToNat ?x] ] =>
assert (Z.to_nat
(Integers.Ptrofs.unsigned
(Integers.Ptrofs.divu
OFFSET
(Integers.Ptrofs.repr 4)))
=
- valueToNat (vdiv (vplus asr # r0 (ZToValue 32 z) ?EQ8) (ZToValue 32 4) ?EQ7))
- as EXPR_OK by admit.
-
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET.
-
- econstructor.
- repeat split; simplify.
- unfold HTL.empty_stack.
- simplify.
- unfold Verilog.merge_arrs.
-
- rewrite AssocMap.gcombine.
- 2: { reflexivity. }
- unfold Verilog.arr_assocmap_set.
- rewrite AssocMap.gss.
- unfold Verilog.merge_arr.
- rewrite AssocMap.gss.
- setoid_rewrite H5.
- reflexivity.
-
- rewrite combine_length.
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- apply list_repeat_len.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- rewrite H4. reflexivity.
-
- intros.
- destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
-
- erewrite Mem.load_store_same.
- 2: { rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite e.
- rewrite Integers.Ptrofs.unsigned_repr.
- exact H1.
- apply Integers.Ptrofs.unsigned_range_2. }
- constructor.
- erewrite combine_lookup_second.
- simpl.
- assert (Ple src (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
- apply H0 in H14.
- destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H14; eauto.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len. auto.
-
- assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption).
- rewrite Z.mul_comm in H14.
- rewrite Z_div_mult in H14; try lia.
- replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H14 by reflexivity.
- rewrite <- PtrofsExtra.divu_unsigned in H14; unfold_constants; try lia.
- rewrite H14. rewrite EXPR_OK.
- rewrite array_get_error_set_bound.
- reflexivity.
- unfold arr_length, arr_repeat. simpl.
- rewrite list_repeat_len. lia.
-
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr.
- simpl.
- destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
- right.
- apply ZExtra.mod_0_bounds; try lia.
- apply ZLib.Z_mod_mult'.
- invert H13.
- rewrite Z2Nat.id in H19; try lia.
- apply Zmult_lt_compat_r with (p := 4) in H19; try lia.
- rewrite ZLib.div_mul_undo in H19; try lia.
- split; try lia.
- apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia.
- }
-
- rewrite <- EXPR_OK.
- rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia).
- destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4).
- apply Z.mul_cancel_r with (p := 4) in e; try lia.
- rewrite ZLib.div_mul_undo in e; try lia.
- rewrite combine_lookup_first.
- eapply H7; eauto.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len. auto.
- rewrite array_gso.
- unfold array_get_error.
- unfold arr_repeat.
- simplify.
- apply list_repeat_lookup.
- lia.
- unfold_constants.
- intro.
- apply Z2Nat.inj_iff in H14; try lia.
- apply Z.div_pos; try lia.
+ valueToNat x)
+ as EXPR_OK by admit
+ end.
+ rewrite <- EXPR_OK.
+
+ specialize (H7 (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu
+ OFFSET
+ (Integers.Ptrofs.repr 4)))).
+ exploit H7; big_tac.
+
+ (** RSBP preservation *)
+ unfold arr_stack_based_pointers in ASBP.
+ specialize (ASBP (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))).
+ exploit ASBP; big_tac.
+ rewrite NORMALISE in H0. rewrite H1 in H0. assumption.
+
+ + invert MARR. crush.
+
+ unfold Op.eval_addressing in H0.
+ destruct (Archi.ptr64) eqn:ARCHI; crush.
+ rewrite ARCHI in H0. crush.
+
+ unfold check_address_parameter_unsigned in *;
+ unfold check_address_parameter_signed in *; crush.
+
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ rewrite ZERO in H1. clear ZERO.
+ rewrite Integers.Ptrofs.add_zero_l in H1.
+
+ remember i0 as OFFSET.
+
+ (** Modular preservation proof *)
+ rename H0 into MOD_PRESERVE.
+
+ (** Read bounds proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as READ_BOUND_HIGH.
+ { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; crush; auto.
+ unfold stack_bounds in BOUNDS.
+ exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET)); big_tac. }
+
+ (** Normalisation proof *)
+ assert (Integers.Ptrofs.repr
+ (4 * Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))) = OFFSET)
+ as NORMALISE.
+ { replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) at 1 by reflexivity.
+ rewrite <- PtrofsExtra.mul_unsigned.
+ apply PtrofsExtra.mul_divu; crush. }
+
+ (** Normalised bounds proof *)
+ assert (0 <=
+ Integers.Ptrofs.unsigned (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4))
+ < (RTL.fn_stacksize f / 4))
+ as NORMALISE_BOUND.
+ { split.
apply Integers.Ptrofs.unsigned_range_2.
+ assert (forall x y, Integers.Ptrofs.divu x y = Integers.Ptrofs.divu x y ) by reflexivity.
+ unfold Integers.Ptrofs.divu at 2 in H0.
+ rewrite H0. clear H0.
+ rewrite Integers.Ptrofs.unsigned_repr; crush.
+ apply Zmult_lt_reg_r with (p := 4); try lia.
+ repeat rewrite ZLib.div_mul_undo; try lia.
+ apply Z.div_pos; small_tac.
+ apply Z.div_le_upper_bound; lia. }
+
+ inversion NORMALISE_BOUND as [ NORMALISE_BOUND_LOW NORMALISE_BOUND_HIGH ];
+ clear NORMALISE_BOUND.
+
+ (** Start of proof proper *)
+ eexists. split.
+ eapply Smallstep.plus_one.
+ eapply HTL.step_module; eauto.
+ apply assumption_32bit.
+ econstructor. econstructor. econstructor. crush.
+ econstructor. econstructor. econstructor. econstructor. crush.
+
+ all: big_tac.
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- unfold arr_stack_based_pointers.
- intros.
- destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
-
- simplify.
- erewrite Mem.load_store_same.
- 2: { rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite e.
- rewrite Integers.Ptrofs.unsigned_repr.
- exact H1.
- apply Integers.Ptrofs.unsigned_range_2. }
- simplify.
- destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor.
- destruct (Archi.ptr64); try discriminate.
- pose proof (RSBP src). rewrite EQ_SRC in H0. assumption.
-
- simpl.
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr.
- simpl.
- destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
- right.
- apply ZExtra.mod_0_bounds; try lia.
- apply ZLib.Z_mod_mult'.
- invert H0.
- apply Zmult_lt_compat_r with (p := 4) in H14; try lia.
- rewrite ZLib.div_mul_undo in H14; try lia.
- split; try lia.
- apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia.
- }
- apply ASBP; assumption.
-
- unfold stack_bounds in *. intros.
- simpl.
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right. right. simpl.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr; simplify; try lia.
- apply ZExtra.mod_0_bounds; simplify; try lia. }
- simplify. split.
- exploit (BOUNDS ptr); try lia. intros. simplify. assumption.
- exploit (BOUNDS ptr v); try lia. intros.
- invert H0.
- match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity.
- assert (Mem.valid_access m AST.Mint32 sp'
- (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
- (Integers.Ptrofs.repr ptr))) Writable).
- { pose proof H1. eapply Mem.store_valid_access_2 in H0.
- exact H0. eapply Mem.store_valid_access_3. eassumption. }
- pose proof (Mem.valid_access_store m AST.Mint32 sp'
- (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
- (Integers.Ptrofs.repr ptr))) v).
- apply X in H0. invert H0. congruence.
-
- + (** Preamble *)
- invert MARR. simplify.
-
- unfold Op.eval_addressing in H0.
- destruct (Archi.ptr64) eqn:ARCHI; simplify.
-
- unfold reg_stack_based_pointers in RSBP.
- pose proof (RSBP r0) as RSBPr0.
- pose proof (RSBP r1) as RSBPr1.
-
- destruct (Registers.Regmap.get r0 rs) eqn:EQr0;
- destruct (Registers.Regmap.get r1 rs) eqn:EQr1; simplify.
-
- rewrite ARCHI in H1. simplify.
- subst.
- clear RSBPr1.
-
- pose proof MASSOC as MASSOC'.
- invert MASSOC'.
- pose proof (H0 r0).
- pose proof (H0 r1).
- assert (HPler0 : Ple r0 (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simplify; eauto).
- assert (HPler1 : Ple r1 (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simpl; auto).
- apply H6 in HPler0.
- apply H8 in HPler1.
- invert HPler0; invert HPler1; try congruence.
- rewrite EQr0 in H10.
- rewrite EQr1 in H12.
- invert H10. invert H12.
- clear H0. clear H6. clear H8.
-
- unfold check_address_parameter_signed in *;
- unfold check_address_parameter_unsigned in *; simplify.
-
- remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (valueToZ asr # r0))
- (Integers.Ptrofs.of_int
- (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z))
- (Integers.Int.repr z0)))) as OFFSET.
-
- (** Modular preservation proof *)
- assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE.
- { rewrite HeqOFFSET.
- apply PtrofsExtra.add_mod; simplify; try lia.
- exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *)
- rewrite Integers.Ptrofs.signed_repr; try assumption.
- admit. (* FIXME: Register bounds. *)
- apply PtrofsExtra.of_int_mod.
- apply IntExtra.add_mod; simplify; try lia.
- exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *)
- apply IntExtra.mul_mod; simplify; try lia.
- exists 1073741824. reflexivity. (* FIXME: This is sadness inducing. *)
- admit. (* FIXME: Register bounds. *)
- rewrite Integers.Int.signed_repr; simplify; try split; try assumption.
- rewrite Integers.Int.signed_repr; simplify; try split; try assumption.
- }
-
- (** Write bounds proof *)
- assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH.
- { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; simplify; auto.
- unfold stack_bounds in BOUNDS.
- exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET) (Registers.Regmap.get src rs)); auto.
- split; try lia; apply Integers.Ptrofs.unsigned_range_2.
- simplify.
- replace (Integers.Ptrofs.repr 0) with (Integers.Ptrofs.zero) by reflexivity.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr. intros. simplify. congruence.
- apply Integers.Ptrofs.unsigned_range_2. }
-
- (** Start of proof proper *)
- eexists. split.
- eapply Smallstep.plus_one.
- eapply HTL.step_module; eauto.
- apply assumption_32bit.
- econstructor. econstructor. econstructor.
- eapply Verilog.stmnt_runp_Vnonblock_arr. simplify.
- econstructor.
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ9]).
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ10]).
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ11]).
- econstructor. econstructor. econstructor. econstructor.
- econstructor.
- eapply Verilog.erun_Vbinop with (EQ := ?[EQ12]).
- econstructor. econstructor. econstructor. econstructor.
- econstructor. econstructor. econstructor. econstructor.
- econstructor. econstructor. econstructor. econstructor.
-
- all: simplify.
-
- (** State Lookup *)
- unfold Verilog.merge_regs.
- simplify.
- unfold_merge.
- apply AssocMap.gss.
-
- (** Match states *)
- rewrite assumption_32bit.
- econstructor; eauto.
-
- (** Match assocmaps *)
- unfold Verilog.merge_regs. simplify. unfold_merge.
- apply regs_lessdef_add_greater. apply greater_than_max_func.
- assumption.
-
- (** States well formed *)
- unfold state_st_wf. inversion 1. simplify.
- unfold Verilog.merge_regs.
- unfold_merge.
- apply AssocMap.gss.
-
- (** Match stacks *)
- admit.
-
- (** Equality proof *)
+ 1: { assert (HPle : Ple dst (RTL.max_reg_function f)).
+ eapply RTL.max_reg_function_def. eassumption. auto.
+ apply ZExtra.Pge_not_eq. apply ZExtra.Ple_Plt_Succ. assumption. }
+
+ 2: { assert (HPle : Ple dst (RTL.max_reg_function f)).
+ eapply RTL.max_reg_function_def. eassumption. auto.
+ apply ZExtra.Pge_not_eq. apply ZExtra.Ple_Plt_Succ. assumption. }
+
+ (** Match assocmaps *)
+ apply regs_lessdef_add_match; big_tac.
+
+ (** Equality proof *)
+ match goal with (* Prevents issues with evars *)
+ | [ |- context [valueToNat ?x] ] =>
assert (Z.to_nat
(Integers.Ptrofs.unsigned
(Integers.Ptrofs.divu
OFFSET
(Integers.Ptrofs.repr 4)))
=
- valueToNat (vdiv
- (vplus (vplus asr # r0 (ZToValue 32 z0) ?EQ11) (vmul asr # r1 (ZToValue 32 z) ?EQ12)
- ?EQ10) (ZToValue 32 4) ?EQ9))
- as EXPR_OK by admit.
-
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET.
-
- econstructor.
- repeat split; simplify.
- unfold HTL.empty_stack.
- simplify.
- unfold Verilog.merge_arrs.
-
- rewrite AssocMap.gcombine.
- 2: { reflexivity. }
- unfold Verilog.arr_assocmap_set.
- rewrite AssocMap.gss.
- unfold Verilog.merge_arr.
- rewrite AssocMap.gss.
- setoid_rewrite H5.
- reflexivity.
-
- rewrite combine_length.
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- apply list_repeat_len.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- rewrite H4. reflexivity.
-
- intros.
- destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
-
- erewrite Mem.load_store_same.
- 2: { rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite e.
- rewrite Integers.Ptrofs.unsigned_repr.
- exact H1.
- apply Integers.Ptrofs.unsigned_range_2. }
- constructor.
- erewrite combine_lookup_second.
- simpl.
- assert (Ple src (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
- apply H0 in H21.
- destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H21; eauto.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len. auto.
-
- assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption).
- rewrite Z.mul_comm in H21.
- rewrite Z_div_mult in H21; try lia.
- replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H21 by reflexivity.
- rewrite <- PtrofsExtra.divu_unsigned in H21; unfold_constants; try lia.
- rewrite H21. rewrite EXPR_OK.
- rewrite array_get_error_set_bound.
- reflexivity.
- unfold arr_length, arr_repeat. simpl.
- rewrite list_repeat_len. lia.
-
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr.
- simpl.
- destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
- right.
- apply ZExtra.mod_0_bounds; try lia.
- apply ZLib.Z_mod_mult'.
- invert H20.
- rewrite Z2Nat.id in H22; try lia.
- apply Zmult_lt_compat_r with (p := 4) in H22; try lia.
- rewrite ZLib.div_mul_undo in H22; try lia.
- split; try lia.
- apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia.
- }
-
- rewrite <- EXPR_OK.
- rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia).
- destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4).
- apply Z.mul_cancel_r with (p := 4) in e; try lia.
- rewrite ZLib.div_mul_undo in e; try lia.
- rewrite combine_lookup_first.
- eapply H7; eauto.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len. auto.
- rewrite array_gso.
- unfold array_get_error.
- unfold arr_repeat.
- simplify.
- apply list_repeat_lookup.
- lia.
- unfold_constants.
- intro.
- apply Z2Nat.inj_iff in H21; try lia.
- apply Z.div_pos; try lia.
- apply Integers.Ptrofs.unsigned_range.
-
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- unfold arr_stack_based_pointers.
- intros.
- destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
-
- simplify.
- erewrite Mem.load_store_same.
- 2: { rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite e.
- rewrite Integers.Ptrofs.unsigned_repr.
- exact H1.
- apply Integers.Ptrofs.unsigned_range_2. }
- simplify.
- destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor.
- destruct (Archi.ptr64); try discriminate.
- pose proof (RSBP src). rewrite EQ_SRC in H0.
- assumption.
-
- simpl.
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr.
- simpl.
- destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
- right.
- apply ZExtra.mod_0_bounds; try lia.
- apply ZLib.Z_mod_mult'.
- invert H0.
- apply Zmult_lt_compat_r with (p := 4) in H21; try lia.
- rewrite ZLib.div_mul_undo in H21; try lia.
- split; try lia.
- apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia.
- }
- apply ASBP; assumption.
-
- unfold stack_bounds in *. intros.
- simpl.
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right. right. simpl.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr; simplify; try lia.
- apply ZExtra.mod_0_bounds; simplify; try lia. }
- simplify. split.
- exploit (BOUNDS ptr); try lia. intros. simplify. assumption.
- exploit (BOUNDS ptr v); try lia. intros.
- invert H0.
- match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity.
- assert (Mem.valid_access m AST.Mint32 sp'
- (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
- (Integers.Ptrofs.repr ptr))) Writable).
- { pose proof H1. eapply Mem.store_valid_access_2 in H0.
- exact H0. eapply Mem.store_valid_access_3. eassumption. }
- pose proof (Mem.valid_access_store m AST.Mint32 sp'
- (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
- (Integers.Ptrofs.repr ptr))) v).
- apply X in H0. invert H0. congruence.
-
- + invert MARR. simplify.
-
- unfold Op.eval_addressing in H0.
- destruct (Archi.ptr64) eqn:ARCHI; simplify.
- rewrite ARCHI in H0. simplify.
-
- unfold check_address_parameter_unsigned in *;
- unfold check_address_parameter_signed in *; simplify.
-
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- rewrite ZERO in H1. clear ZERO.
- rewrite Integers.Ptrofs.add_zero_l in H1.
-
- remember i0 as OFFSET.
-
- (** Modular preservation proof *)
- rename H0 into MOD_PRESERVE.
-
- (** Write bounds proof *)
- assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH.
- { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; simplify; auto.
- unfold stack_bounds in BOUNDS.
- exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET) (Registers.Regmap.get src rs)); auto.
- simplify.
- replace (Integers.Ptrofs.repr 0) with (Integers.Ptrofs.zero) by reflexivity.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr. intros. simplify. congruence.
- apply Integers.Ptrofs.unsigned_range_2. }
-
- (** Start of proof proper *)
- eexists. split.
- eapply Smallstep.plus_one.
- eapply HTL.step_module; eauto.
- apply assumption_32bit.
- econstructor. econstructor. econstructor.
- eapply Verilog.stmnt_runp_Vnonblock_arr. simplify.
- econstructor. econstructor. econstructor. econstructor.
-
- all: simplify.
-
- (** State Lookup *)
- unfold Verilog.merge_regs.
- simplify.
- unfold_merge.
- apply AssocMap.gss.
-
- (** Match states *)
- rewrite assumption_32bit.
- econstructor; eauto.
-
- (** Match assocmaps *)
- unfold Verilog.merge_regs. simplify. unfold_merge.
- apply regs_lessdef_add_greater. apply greater_than_max_func.
- assumption.
-
- (** States well formed *)
- unfold state_st_wf. inversion 1. simplify.
- unfold Verilog.merge_regs.
- unfold_merge.
- apply AssocMap.gss.
-
- (** Match stacks *)
- admit.
-
- (** Equality proof *)
+ valueToNat x)
+ as EXPR_OK by admit
+ end.
+ rewrite <- EXPR_OK.
+
+ specialize (H7 (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu
+ OFFSET
+ (Integers.Ptrofs.repr 4)))).
+ exploit H7; big_tac.
+
+ (** RSBP preservation *)
+ unfold arr_stack_based_pointers in ASBP.
+ specialize (ASBP (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu OFFSET (Integers.Ptrofs.repr 4)))).
+ exploit ASBP; big_tac.
+ rewrite NORMALISE in H0. rewrite H1 in H0. assumption.*)
+ Admitted.
+ Hint Resolve transl_iload_correct : htlproof.
+
+ Lemma transl_istore_correct:
+ forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive)
+ (rs : Registers.Regmap.t Values.val) (m : mem) (chunk : AST.memory_chunk)
+ (addr : Op.addressing) (args : list Registers.reg) (src : Registers.reg)
+ (pc' : RTL.node) (a : Values.val) (m' : mem),
+ (RTL.fn_code f) ! pc = Some (RTL.Istore chunk addr args src pc') ->
+ Op.eval_addressing ge sp addr (map (fun r : positive => Registers.Regmap.get r rs) args) = Some a ->
+ Mem.storev chunk m a (Registers.Regmap.get src rs) = Some m' ->
+ forall R1 : HTL.state,
+ match_states (RTL.State s f sp pc rs m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m') R2.
+ Proof.
+(* intros s f sp pc rs m chunk addr args src pc' a m' H H0 H1 R1 MSTATES.
+ inv_state. inv_arr_access.
+
+ + (** Preamble *)
+ invert MARR. crush.
+
+ unfold Op.eval_addressing in H0.
+ destruct (Archi.ptr64) eqn:ARCHI; crush.
+
+ unfold reg_stack_based_pointers in RSBP.
+ pose proof (RSBP r0) as RSBPr0.
+
+ destruct (Registers.Regmap.get r0 rs) eqn:EQr0; crush.
+
+ rewrite ARCHI in H1. crush.
+ subst.
+
+ pose proof MASSOC as MASSOC'.
+ invert MASSOC'.
+ pose proof (H0 r0).
+ assert (HPler0 : Ple r0 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; crush; eauto).
+ apply H6 in HPler0.
+ invert HPler0; try congruence.
+ rewrite EQr0 in H8.
+ invert H8.
+ clear H0. clear H6.
+
+ unfold check_address_parameter_unsigned in *;
+ unfold check_address_parameter_signed in *; crush.
+
+ remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ asr # r0))
+ (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET.
+
+ (** Modular preservation proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE.
+ { rewrite HeqOFFSET.
+ apply PtrofsExtra.add_mod; crush; try lia.
+ rewrite Integers.Ptrofs.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush.
+ apply PtrofsExtra.of_int_mod.
+ rewrite Integers.Int.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush. }
+
+ (** Write bounds proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH.
+ { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; crush; auto.
+ unfold stack_bounds in BOUNDS.
+ exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET) (Registers.Regmap.get src rs)); big_tac.
+ apply Integers.Ptrofs.unsigned_range_2. }
+
+ (** Start of proof proper *)
+ eexists. split.
+ eapply Smallstep.plus_one.
+ eapply HTL.step_module; eauto.
+ apply assumption_32bit.
+ econstructor. econstructor. econstructor.
+ eapply Verilog.stmnt_runp_Vnonblock_arr. crush.
+ econstructor.
+ eapply Verilog.erun_Vbinop with (EQ := ?[EQ7]).
+ eapply Verilog.erun_Vbinop with (EQ := ?[EQ8]).
+ econstructor.
+ econstructor.
+ econstructor. econstructor. econstructor. econstructor.
+ econstructor. econstructor. econstructor. econstructor.
+
+ all: crush.
+
+ (** State Lookup *)
+ unfold Verilog.merge_regs.
+ crush.
+ unfold_merge.
+ apply AssocMap.gss.
+
+ (** Match states *)
+ rewrite assumption_32bit.
+ econstructor; eauto.
+
+ (** Match assocmaps *)
+ unfold Verilog.merge_regs. crush. unfold_merge.
+ apply regs_lessdef_add_greater.
+ unfold Plt; lia.
+ assumption.
+
+ (** States well formed *)
+ unfold state_st_wf. inversion 1. crush.
+ unfold Verilog.merge_regs.
+ unfold_merge.
+ apply AssocMap.gss.
+
+ (** Equality proof *)
+ match goal with
+ | [ |- context [valueToNat ?x] ] =>
assert (Z.to_nat
(Integers.Ptrofs.unsigned
(Integers.Ptrofs.divu
OFFSET
(Integers.Ptrofs.repr 4)))
=
- valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned OFFSET / 4)))
- as EXPR_OK by admit.
-
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET.
-
- econstructor.
- repeat split; simplify.
- unfold HTL.empty_stack.
- simplify.
- unfold Verilog.merge_arrs.
-
- rewrite AssocMap.gcombine.
- 2: { reflexivity. }
- unfold Verilog.arr_assocmap_set.
- rewrite AssocMap.gss.
- unfold Verilog.merge_arr.
- rewrite AssocMap.gss.
- setoid_rewrite H5.
- reflexivity.
-
- rewrite combine_length.
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- apply list_repeat_len.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- rewrite H4. reflexivity.
-
- intros.
- destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
-
- erewrite Mem.load_store_same.
- 2: { rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite e.
- rewrite Integers.Ptrofs.unsigned_repr.
- exact H1.
- apply Integers.Ptrofs.unsigned_range_2. }
- constructor.
- erewrite combine_lookup_second.
- simpl.
- assert (Ple src (RTL.max_reg_function f))
- by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
- apply H0 in H10.
- destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H10; eauto.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len. auto.
-
- assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption).
- rewrite Z.mul_comm in H10.
- rewrite Z_div_mult in H10; try lia.
- replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H10 by reflexivity.
- rewrite <- PtrofsExtra.divu_unsigned in H10; unfold_constants; try lia.
- rewrite H10. rewrite EXPR_OK.
- rewrite array_get_error_set_bound.
- reflexivity.
- unfold arr_length, arr_repeat. simpl.
- rewrite list_repeat_len. lia.
-
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr.
- simpl.
- destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
- right.
- apply ZExtra.mod_0_bounds; try lia.
- apply ZLib.Z_mod_mult'.
- invert H8.
- rewrite Z2Nat.id in H12; try lia.
- apply Zmult_lt_compat_r with (p := 4) in H12; try lia.
- rewrite ZLib.div_mul_undo in H12; try lia.
- split; try lia.
- apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia.
- }
-
- rewrite <- EXPR_OK.
- rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia).
- destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4).
- apply Z.mul_cancel_r with (p := 4) in e; try lia.
- rewrite ZLib.div_mul_undo in e; try lia.
- rewrite combine_lookup_first.
- eapply H7; eauto.
-
- rewrite <- array_set_len.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len. auto.
- rewrite array_gso.
- unfold array_get_error.
- unfold arr_repeat.
- simplify.
- apply list_repeat_lookup.
- lia.
- unfold_constants.
- intro.
- apply Z2Nat.inj_iff in H10; try lia.
- apply Z.div_pos; try lia.
- apply Integers.Ptrofs.unsigned_range.
-
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- unfold arr_stack_based_pointers.
- intros.
- destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
-
- simplify.
- erewrite Mem.load_store_same.
- 2: { rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite e.
- rewrite Integers.Ptrofs.unsigned_repr.
- exact H1.
- apply Integers.Ptrofs.unsigned_range_2. }
- simplify.
- destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor.
- destruct (Archi.ptr64); try discriminate.
- pose proof (RSBP src). rewrite EQ_SRC in H0.
- assumption.
-
- simpl.
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr.
- simpl.
- destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
- right.
- apply ZExtra.mod_0_bounds; try lia.
- apply ZLib.Z_mod_mult'.
- invert H0.
- apply Zmult_lt_compat_r with (p := 4) in H10; try lia.
- rewrite ZLib.div_mul_undo in H10; try lia.
- split; try lia.
- apply Z.le_trans with (m := RTL.fn_stacksize f); simplify; lia.
- }
- apply ASBP; assumption.
-
- unfold stack_bounds in *. intros.
- simpl.
- assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
- erewrite Mem.load_store_other with (m1 := m).
- 2: { exact H1. }
- 2: { right. right. simpl.
- rewrite ZERO.
- rewrite Integers.Ptrofs.add_zero_l.
- rewrite Integers.Ptrofs.unsigned_repr; simplify; try lia.
- apply ZExtra.mod_0_bounds; simplify; try lia. }
- simplify. split.
- exploit (BOUNDS ptr); try lia. intros. simplify. assumption.
- exploit (BOUNDS ptr v); try lia. intros.
- invert H0.
- match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity.
- assert (Mem.valid_access m AST.Mint32 sp'
- (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
- (Integers.Ptrofs.repr ptr))) Writable).
- { pose proof H1. eapply Mem.store_valid_access_2 in H0.
- exact H0. eapply Mem.store_valid_access_3. eassumption. }
- pose proof (Mem.valid_access_store m AST.Mint32 sp'
- (Integers.Ptrofs.unsigned
- (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
- (Integers.Ptrofs.repr ptr))) v).
- apply X in H0. invert H0. congruence.
-
- - eexists. split. apply Smallstep.plus_one.
- eapply HTL.step_module; eauto.
- apply assumption_32bit.
- eapply Verilog.stmnt_runp_Vnonblock_reg with
- (rhsval := if b then posToValue 32 ifso else posToValue 32 ifnot).
+ valueToNat x)
+ as EXPR_OK by admit
+ end.
+
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET.
+
+ econstructor.
+ repeat split; crush.
+ unfold HTL.empty_stack.
+ crush.
+ unfold Verilog.merge_arrs.
+
+ rewrite AssocMap.gcombine.
+ 2: { reflexivity. }
+ unfold Verilog.arr_assocmap_set.
+ rewrite AssocMap.gss.
+ unfold Verilog.merge_arr.
+ rewrite AssocMap.gss.
+ setoid_rewrite H5.
+ reflexivity.
+
+ rewrite combine_length.
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ apply list_repeat_len.
+
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ rewrite list_repeat_len.
+ rewrite H4. reflexivity.
+
+ remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ asr # r0))
+ (Integers.Ptrofs.of_int (Integers.Int.repr z))) as OFFSET.
+
+ destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
+ erewrite Mem.load_store_same.
+ 2: { rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite e.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ exact H1.
+ apply Integers.Ptrofs.unsigned_range_2. }
constructor.
+ erewrite combine_lookup_second.
+ simpl.
+ assert (Ple src (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ apply H0 in H13.
+ destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H13; eauto.
+
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ rewrite list_repeat_len. auto.
+
+ assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption).
+ rewrite Z.mul_comm in H13.
+ rewrite Z_div_mult in H13; try lia.
+ replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H13 by reflexivity.
+ rewrite <- PtrofsExtra.divu_unsigned in H13; unfold_constants; try lia.
+ rewrite H13. rewrite EXPR_OK.
+ rewrite array_get_error_set_bound.
+ reflexivity.
+ unfold arr_length, arr_repeat. simpl.
+ rewrite list_repeat_len. lia.
+
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ simpl.
+ destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
+ right.
+ apply ZExtra.mod_0_bounds; try lia.
+ apply ZLib.Z_mod_mult'.
+ rewrite Z2Nat.id in H15; try lia.
+ apply Zmult_lt_compat_r with (p := 4) in H15; try lia.
+ rewrite ZLib.div_mul_undo in H15; try lia.
+ split; try lia.
+ apply Z.le_trans with (m := RTL.fn_stacksize f); crush; lia.
+ }
+
+ rewrite <- EXPR_OK.
+ rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia).
+ destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4).
+ apply Z.mul_cancel_r with (p := 4) in e; try lia.
+ rewrite ZLib.div_mul_undo in e; try lia.
+ rewrite combine_lookup_first.
+ eapply H7; eauto.
+
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ rewrite list_repeat_len. auto.
+ rewrite array_gso.
+ unfold array_get_error.
+ unfold arr_repeat.
+ crush.
+ apply list_repeat_lookup.
+ lia.
+ unfold_constants.
+ intro.
+ apply Z2Nat.inj_iff in H13; try lia.
+ apply Z.div_pos; try lia.
+ apply Integers.Ptrofs.unsigned_range.
+
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ unfold arr_stack_based_pointers.
+ intros.
+ destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
+
+ crush.
+ erewrite Mem.load_store_same.
+ 2: { rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite e.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ exact H1.
+ apply Integers.Ptrofs.unsigned_range_2. }
+ crush.
+ destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor.
+ destruct (Archi.ptr64); try discriminate.
+ pose proof (RSBP src). rewrite EQ_SRC in H0.
+ assumption.
simpl.
- destruct b.
- eapply Verilog.erun_Vternary_true.
- eapply eval_cond_correct; eauto.
- constructor.
- apply boolToValue_ValueToBool.
- eapply Verilog.erun_Vternary_false.
- eapply eval_cond_correct; eauto.
- constructor.
- apply boolToValue_ValueToBool.
- constructor.
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ simpl.
+ destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
+ right.
+ apply ZExtra.mod_0_bounds; try lia.
+ apply ZLib.Z_mod_mult'.
+ invert H0.
+ apply Zmult_lt_compat_r with (p := 4) in H14; try lia.
+ rewrite ZLib.div_mul_undo in H14; try lia.
+ split; try lia.
+ apply Z.le_trans with (m := RTL.fn_stacksize f); crush; lia.
+ }
+ apply ASBP; assumption.
+
+ unfold stack_bounds in *. intros.
+ simpl.
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right. right. simpl.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr; crush; try lia.
+ apply ZExtra.mod_0_bounds; crush; try lia. }
+ crush.
+ exploit (BOUNDS ptr); try lia. intros. crush.
+ exploit (BOUNDS ptr v); try lia. intros.
+ invert H0.
+ match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity.
+ assert (Mem.valid_access m AST.Mint32 sp'
+ (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
+ (Integers.Ptrofs.repr ptr))) Writable).
+ { pose proof H1. eapply Mem.store_valid_access_2 in H0.
+ exact H0. eapply Mem.store_valid_access_3. eassumption. }
+ pose proof (Mem.valid_access_store m AST.Mint32 sp'
+ (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
+ (Integers.Ptrofs.repr ptr))) v).
+ apply X in H0. invert H0. congruence.
+
+ + (** Preamble *)
+ invert MARR. crush.
+
+ unfold Op.eval_addressing in H0.
+ destruct (Archi.ptr64) eqn:ARCHI; crush.
+
+ unfold reg_stack_based_pointers in RSBP.
+ pose proof (RSBP r0) as RSBPr0.
+ pose proof (RSBP r1) as RSBPr1.
+
+ destruct (Registers.Regmap.get r0 rs) eqn:EQr0;
+ destruct (Registers.Regmap.get r1 rs) eqn:EQr1; crush.
+
+ rewrite ARCHI in H1. crush.
+ subst.
+ clear RSBPr1.
+
+ pose proof MASSOC as MASSOC'.
+ invert MASSOC'.
+ pose proof (H0 r0).
+ pose proof (H0 r1).
+ assert (HPler0 : Ple r0 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; crush; eauto).
+ assert (HPler1 : Ple r1 (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; simpl; auto).
+ apply H6 in HPler0.
+ apply H8 in HPler1.
+ invert HPler0; invert HPler1; try congruence.
+ rewrite EQr0 in H9.
+ rewrite EQr1 in H11.
+ invert H9. invert H11.
+ clear H0. clear H6. clear H8.
+
+ unfold check_address_parameter_signed in *;
+ unfold check_address_parameter_unsigned in *; crush.
+
+ remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ asr # r0))
+ (Integers.Ptrofs.of_int
+ (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z))
+ (Integers.Int.repr z0)))) as OFFSET.
+
+ (** Modular preservation proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET mod 4 = 0) as MOD_PRESERVE.
+ { rewrite HeqOFFSET.
+ apply PtrofsExtra.add_mod; crush; try lia.
+ rewrite Integers.Ptrofs.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush.
+ apply PtrofsExtra.of_int_mod.
+ apply IntExtra.add_mod; crush.
+ apply IntExtra.mul_mod2; crush.
+ rewrite Integers.Int.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush.
+ rewrite Integers.Int.unsigned_repr_eq.
+ rewrite <- Zmod_div_mod; crush. }
+
+ (** Write bounds proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH.
+ { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; crush; auto.
+ unfold stack_bounds in BOUNDS.
+ exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET) (Registers.Regmap.get src rs)); auto.
+ split; try lia; apply Integers.Ptrofs.unsigned_range_2.
+ small_tac. }
+
+ (** Start of proof proper *)
+ eexists. split.
+ eapply Smallstep.plus_one.
+ eapply HTL.step_module; eauto.
+ apply assumption_32bit.
+ econstructor. econstructor. econstructor.
+ eapply Verilog.stmnt_runp_Vnonblock_arr. crush.
+ econstructor.
+ eapply Verilog.erun_Vbinop with (EQ := ?[EQ9]).
+ eapply Verilog.erun_Vbinop with (EQ := ?[EQ10]).
+ eapply Verilog.erun_Vbinop with (EQ := ?[EQ11]).
+ econstructor. econstructor. econstructor. econstructor.
+ econstructor.
+ eapply Verilog.erun_Vbinop with (EQ := ?[EQ12]).
+ econstructor. econstructor. econstructor. econstructor.
+ econstructor. econstructor. econstructor. econstructor.
+ econstructor. econstructor. econstructor. econstructor.
+
+ all: crush.
+
+ (** State Lookup *)
unfold Verilog.merge_regs.
+ crush.
unfold_merge.
apply AssocMap.gss.
- destruct b.
+ (** Match states *)
rewrite assumption_32bit.
- simplify.
- apply match_state with (sp' := sp'); eauto.
- unfold Verilog.merge_regs.
- unfold_merge.
- apply regs_lessdef_add_greater. apply greater_than_max_func.
+ econstructor; eauto.
+
+ (** Match assocmaps *)
+ unfold Verilog.merge_regs. crush. unfold_merge.
+ apply regs_lessdef_add_greater.
+ unfold Plt; lia.
assumption.
- unfold state_st_wf. intros.
- invert H3.
- unfold Verilog.merge_regs. unfold_merge.
+ (** States well formed *)
+ unfold state_st_wf. inversion 1. crush.
+ unfold Verilog.merge_regs.
+ unfold_merge.
apply AssocMap.gss.
- (** Match arrays *)
- invert MARR. simplify.
+ (** Equality proof *)
+ assert (Z.to_nat
+ (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu
+ OFFSET
+ (Integers.Ptrofs.repr 4)))
+ =
+ valueToNat (vdiv
+ (vplus (vplus asr # r0 (ZToValue 32 z0) ?EQ11) (vmul asr # r1 (ZToValue 32 z) ?EQ12)
+ ?EQ10) (ZToValue 32 4) ?EQ9))
+ as EXPR_OK by admit.
+
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET.
+
econstructor.
- repeat split; simplify.
+ repeat split; crush.
unfold HTL.empty_stack.
- simplify.
+ crush.
unfold Verilog.merge_arrs.
rewrite AssocMap.gcombine.
2: { reflexivity. }
+ unfold Verilog.arr_assocmap_set.
rewrite AssocMap.gss.
unfold Verilog.merge_arr.
- setoid_rewrite H4.
+ rewrite AssocMap.gss.
+ setoid_rewrite H5.
reflexivity.
rewrite combine_length.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- reflexivity.
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ apply list_repeat_len.
- unfold arr_repeat. simplify.
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
rewrite list_repeat_len.
- congruence.
-
+ rewrite H4. reflexivity.
+
+ remember (Integers.Ptrofs.add (Integers.Ptrofs.repr (uvalueToZ asr # r0))
+ (Integers.Ptrofs.of_int
+ (Integers.Int.add (Integers.Int.mul (valueToInt asr # r1) (Integers.Int.repr z))
+ (Integers.Int.repr z0)))) as OFFSET.
+ destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
+
+ erewrite Mem.load_store_same.
+ 2: { rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite e.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ exact H1.
+ apply Integers.Ptrofs.unsigned_range_2. }
+ constructor.
+ erewrite combine_lookup_second.
+ simpl.
+ assert (Ple src (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ apply H0 in H16.
+ destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H16; eauto.
+
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ rewrite list_repeat_len. auto.
+
+ assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption).
+ rewrite Z.mul_comm in H16.
+ rewrite Z_div_mult in H16; try lia.
+ replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H16 by reflexivity.
+ rewrite <- PtrofsExtra.divu_unsigned in H16; unfold_constants; try lia.
+ rewrite H16. rewrite EXPR_OK.
+ rewrite array_get_error_set_bound.
+ reflexivity.
+ unfold arr_length, arr_repeat. simpl.
+ rewrite list_repeat_len. lia.
+
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ simpl.
+ destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
+ right.
+ apply ZExtra.mod_0_bounds; try lia.
+ apply ZLib.Z_mod_mult'.
+ rewrite Z2Nat.id in H18; try lia.
+ apply Zmult_lt_compat_r with (p := 4) in H18; try lia.
+ rewrite ZLib.div_mul_undo in H18; try lia.
+ split; try lia.
+ apply Z.le_trans with (m := RTL.fn_stacksize f); crush; lia.
+ }
+
+ rewrite <- EXPR_OK.
+ rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia).
+ destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4).
+ apply Z.mul_cancel_r with (p := 4) in e; try lia.
+ rewrite ZLib.div_mul_undo in e; try lia.
+ rewrite combine_lookup_first.
+ eapply H7; eauto.
+
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ rewrite list_repeat_len. auto.
+ rewrite array_gso.
+ unfold array_get_error.
+ unfold arr_repeat.
+ crush.
+ apply list_repeat_lookup.
+ lia.
+ unfold_constants.
+ intro.
+ apply Z2Nat.inj_iff in H16; try lia.
+ apply Z.div_pos; try lia.
+ apply Integers.Ptrofs.unsigned_range.
+
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ unfold arr_stack_based_pointers.
intros.
- erewrite array_get_error_equal.
- eauto. apply combine_none.
+ destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
+
+ crush.
+ erewrite Mem.load_store_same.
+ 2: { rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite e.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ exact H1.
+ apply Integers.Ptrofs.unsigned_range_2. }
+ crush.
+ destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor.
+ destruct (Archi.ptr64); try discriminate.
+ pose proof (RSBP src). rewrite EQ_SRC in H0.
assumption.
+ simpl.
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ simpl.
+ destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
+ right.
+ apply ZExtra.mod_0_bounds; try lia.
+ apply ZLib.Z_mod_mult'.
+ invert H0.
+ apply Zmult_lt_compat_r with (p := 4) in H17; try lia.
+ rewrite ZLib.div_mul_undo in H17; try lia.
+ split; try lia.
+ apply Z.le_trans with (m := RTL.fn_stacksize f); crush; lia.
+ }
+ apply ASBP; assumption.
+
+ unfold stack_bounds in *. intros.
+ simpl.
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right. right. simpl.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr; crush; try lia.
+ apply ZExtra.mod_0_bounds; crush; try lia. }
+ crush.
+ exploit (BOUNDS ptr); try lia. intros. crush.
+ exploit (BOUNDS ptr v); try lia. intros.
+ invert H0.
+ match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity.
+ assert (Mem.valid_access m AST.Mint32 sp'
+ (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
+ (Integers.Ptrofs.repr ptr))) Writable).
+ { pose proof H1. eapply Mem.store_valid_access_2 in H0.
+ exact H0. eapply Mem.store_valid_access_3. eassumption. }
+ pose proof (Mem.valid_access_store m AST.Mint32 sp'
+ (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
+ (Integers.Ptrofs.repr ptr))) v).
+ apply X in H0. invert H0. congruence.
+
+ + invert MARR. crush.
+
+ unfold Op.eval_addressing in H0.
+ destruct (Archi.ptr64) eqn:ARCHI; crush.
+ rewrite ARCHI in H0. crush.
+
+ unfold check_address_parameter_unsigned in *;
+ unfold check_address_parameter_signed in *; crush.
+
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ rewrite ZERO in H1. clear ZERO.
+ rewrite Integers.Ptrofs.add_zero_l in H1.
+
+ remember i0 as OFFSET.
+
+ (** Modular preservation proof *)
+ rename H0 into MOD_PRESERVE.
+
+ (** Write bounds proof *)
+ assert (Integers.Ptrofs.unsigned OFFSET < f.(RTL.fn_stacksize)) as WRITE_BOUND_HIGH.
+ { destruct (Integers.Ptrofs.unsigned OFFSET <? f.(RTL.fn_stacksize)) eqn:EQ; crush; auto.
+ unfold stack_bounds in BOUNDS.
+ exploit (BOUNDS (Integers.Ptrofs.unsigned OFFSET) (Registers.Regmap.get src rs)); auto.
+ crush.
+ replace (Integers.Ptrofs.repr 0) with (Integers.Ptrofs.zero) by reflexivity.
+ small_tac. }
+
+ (** Start of proof proper *)
+ eexists. split.
+ eapply Smallstep.plus_one.
+ eapply HTL.step_module; eauto.
+ apply assumption_32bit.
+ econstructor. econstructor. econstructor.
+ eapply Verilog.stmnt_runp_Vnonblock_arr. crush.
+ econstructor. econstructor. econstructor. econstructor.
+
+ all: crush.
+
+ (** State Lookup *)
+ unfold Verilog.merge_regs.
+ crush.
+ unfold_merge.
+ apply AssocMap.gss.
+
+ (** Match states *)
rewrite assumption_32bit.
- apply match_state with (sp' := sp'); eauto.
- unfold Verilog.merge_regs. unfold_merge.
- apply regs_lessdef_add_greater. apply greater_than_max_func.
+ econstructor; eauto.
+
+ (** Match assocmaps *)
+ unfold Verilog.merge_regs. crush. unfold_merge.
+ apply regs_lessdef_add_greater.
+ unfold Plt; lia.
assumption.
- unfold state_st_wf. intros.
- invert H1.
- unfold Verilog.merge_regs. unfold_merge.
+ (** States well formed *)
+ unfold state_st_wf. inversion 1. crush.
+ unfold Verilog.merge_regs.
+ unfold_merge.
apply AssocMap.gss.
- (** Match arrays *)
- invert MARR. simplify.
+ (** Equality proof *)
+ assert (Z.to_nat
+ (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.divu
+ OFFSET
+ (Integers.Ptrofs.repr 4)))
+ =
+ valueToNat (ZToValue 32 (Integers.Ptrofs.unsigned OFFSET / 4)))
+ as EXPR_OK by admit.
+
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ inversion MASSOC; revert HeqOFFSET; subst; clear MASSOC; intros HeqOFFSET.
+
econstructor.
- repeat split; simplify.
+ repeat split; crush.
unfold HTL.empty_stack.
- simplify.
+ crush.
unfold Verilog.merge_arrs.
rewrite AssocMap.gcombine.
2: { reflexivity. }
+ unfold Verilog.arr_assocmap_set.
rewrite AssocMap.gss.
unfold Verilog.merge_arr.
- setoid_rewrite H2.
+ rewrite AssocMap.gss.
+ setoid_rewrite H5.
reflexivity.
rewrite combine_length.
- unfold arr_repeat. simplify.
- rewrite list_repeat_len.
- reflexivity.
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ apply list_repeat_len.
- unfold arr_repeat. simplify.
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
rewrite list_repeat_len.
- congruence.
-
+ rewrite H4. reflexivity.
+
+ remember i0 as OFFSET.
+ destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
+
+ erewrite Mem.load_store_same.
+ 2: { rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite e.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ exact H1.
+ apply Integers.Ptrofs.unsigned_range_2. }
+ constructor.
+ erewrite combine_lookup_second.
+ simpl.
+ assert (Ple src (RTL.max_reg_function f))
+ by (eapply RTL.max_reg_function_use; eauto; simpl; auto);
+ apply H0 in H8.
+ destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; constructor; invert H8; eauto.
+
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ rewrite list_repeat_len. auto.
+
+ assert (4 * ptr / 4 = Integers.Ptrofs.unsigned OFFSET / 4) by (f_equal; assumption).
+ rewrite Z.mul_comm in H8.
+ rewrite Z_div_mult in H8; try lia.
+ replace 4 with (Integers.Ptrofs.unsigned (Integers.Ptrofs.repr 4)) in H8 by reflexivity.
+ rewrite <- PtrofsExtra.divu_unsigned in H8; unfold_constants; try lia.
+ rewrite H8. rewrite EXPR_OK.
+ rewrite array_get_error_set_bound.
+ reflexivity.
+ unfold arr_length, arr_repeat. simpl.
+ rewrite list_repeat_len. lia.
+
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ simpl.
+ destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
+ right.
+ apply ZExtra.mod_0_bounds; try lia.
+ apply ZLib.Z_mod_mult'.
+ rewrite Z2Nat.id in H11; try lia.
+ apply Zmult_lt_compat_r with (p := 4) in H11; try lia.
+ rewrite ZLib.div_mul_undo in H11; try lia.
+ split; try lia.
+ apply Z.le_trans with (m := RTL.fn_stacksize f); crush; lia.
+ }
+
+ rewrite <- EXPR_OK.
+ rewrite PtrofsExtra.divu_unsigned; auto; try (unfold_constants; lia).
+ destruct (ptr ==Z Integers.Ptrofs.unsigned OFFSET / 4).
+ apply Z.mul_cancel_r with (p := 4) in e; try lia.
+ rewrite ZLib.div_mul_undo in e; try lia.
+ rewrite combine_lookup_first.
+ eapply H7; eauto.
+
+ rewrite <- array_set_len.
+ unfold arr_repeat. crush.
+ rewrite list_repeat_len. auto.
+ rewrite array_gso.
+ unfold array_get_error.
+ unfold arr_repeat.
+ crush.
+ apply list_repeat_lookup.
+ lia.
+ unfold_constants.
+ intro.
+ apply Z2Nat.inj_iff in H8; try lia.
+ apply Z.div_pos; try lia.
+ apply Integers.Ptrofs.unsigned_range.
+
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ unfold arr_stack_based_pointers.
intros.
- erewrite array_get_error_equal.
- eauto. apply combine_none.
+ destruct (4 * ptr ==Z Integers.Ptrofs.unsigned OFFSET).
+
+ crush.
+ erewrite Mem.load_store_same.
+ 2: { rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite e.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ exact H1.
+ apply Integers.Ptrofs.unsigned_range_2. }
+ crush.
+ destruct (Registers.Regmap.get src rs) eqn:EQ_SRC; try constructor.
+ destruct (Archi.ptr64); try discriminate.
+ pose proof (RSBP src). rewrite EQ_SRC in H0.
assumption.
- - admit.
+ simpl.
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr.
+ simpl.
+ destruct (Z_le_gt_dec (4 * ptr + 4) (Integers.Ptrofs.unsigned OFFSET)); eauto.
+ right.
+ apply ZExtra.mod_0_bounds; try lia.
+ apply ZLib.Z_mod_mult'.
+ invert H0.
+ apply Zmult_lt_compat_r with (p := 4) in H9; try lia.
+ rewrite ZLib.div_mul_undo in H9; try lia.
+ split; try lia.
+ apply Z.le_trans with (m := RTL.fn_stacksize f); crush; lia.
+ }
+ apply ASBP; assumption.
+
+ unfold stack_bounds in *. intros.
+ simpl.
+ assert (Integers.Ptrofs.repr 0 = Integers.Ptrofs.zero) as ZERO by reflexivity.
+ erewrite Mem.load_store_other with (m1 := m).
+ 2: { exact H1. }
+ 2: { right. right. simpl.
+ rewrite ZERO.
+ rewrite Integers.Ptrofs.add_zero_l.
+ rewrite Integers.Ptrofs.unsigned_repr; crush; try lia.
+ apply ZExtra.mod_0_bounds; crush; try lia. }
+ crush.
+ exploit (BOUNDS ptr); try lia. intros. crush.
+ exploit (BOUNDS ptr v); try lia. intros.
+ invert H0.
+ match goal with | |- ?x = _ => destruct x eqn:EQ end; try reflexivity.
+ assert (Mem.valid_access m AST.Mint32 sp'
+ (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
+ (Integers.Ptrofs.repr ptr))) Writable).
+ { pose proof H1. eapply Mem.store_valid_access_2 in H0.
+ exact H0. eapply Mem.store_valid_access_3. eassumption. }
+ pose proof (Mem.valid_access_store m AST.Mint32 sp'
+ (Integers.Ptrofs.unsigned
+ (Integers.Ptrofs.add (Integers.Ptrofs.repr 0)
+ (Integers.Ptrofs.repr ptr))) v).
+ apply X in H0. invert H0. congruence.*)
+ Admitted.
+ Hint Resolve transl_istore_correct : htlproof.
+
+ Lemma transl_icond_correct:
+ forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive)
+ (rs : Registers.Regmap.t Values.val) (m : mem) (cond : Op.condition) (args : list Registers.reg)
+ (ifso ifnot : RTL.node) (b : bool) (pc' : RTL.node),
+ (RTL.fn_code f) ! pc = Some (RTL.Icond cond args ifso ifnot) ->
+ Op.eval_condition cond (map (fun r : positive => Registers.Regmap.get r rs) args) m = Some b ->
+ pc' = (if b then ifso else ifnot) ->
+ forall R1 : HTL.state,
+ match_states (RTL.State s f sp pc rs m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m) R2.
+ Proof.
+ intros s f sp pc rs m cond args ifso ifnot b pc' H H0 H1 R1 MSTATE.
+ inv_state.
+
+ eexists. split. apply Smallstep.plus_one.
+ eapply HTL.step_module; eauto.
+ inv CONST; assumption.
+ inv CONST; assumption.
+(* eapply Verilog.stmnt_runp_Vnonblock_reg with
+ (rhsval := if b then posToValue 32 ifso else posToValue 32 ifnot).
+ constructor.
- - (* Return *)
- econstructor. split.
+ simpl.
+ destruct b.
+ eapply Verilog.erun_Vternary_true.
+ eapply eval_cond_correct; eauto.
+ constructor.
+ apply boolToValue_ValueToBool.
+ eapply Verilog.erun_Vternary_false.
+ eapply eval_cond_correct; eauto.
+ constructor.
+ apply boolToValue_ValueToBool.
+ constructor.
+
+ big_tac.
+
+ invert MARR.
+ destruct b; rewrite assumption_32bit; big_tac.
+
+ Unshelve.
+ constructor.
+ Qed.*)
+ Admitted.
+ Hint Resolve transl_icond_correct : htlproof.
+
+ Lemma transl_ijumptable_correct:
+ forall (s : list RTL.stackframe) (f : RTL.function) (sp : Values.val) (pc : positive)
+ (rs : Registers.Regmap.t Values.val) (m : mem) (arg : Registers.reg) (tbl : list RTL.node)
+ (n : Integers.Int.int) (pc' : RTL.node),
+ (RTL.fn_code f) ! pc = Some (RTL.Ijumptable arg tbl) ->
+ Registers.Regmap.get arg rs = Values.Vint n ->
+ list_nth_z tbl (Integers.Int.unsigned n) = Some pc' ->
+ forall R1 : HTL.state,
+ match_states (RTL.State s f sp pc rs m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m) R2.
+ Proof.
+ intros s f sp pc rs m arg tbl n pc' H H0 H1 R1 MSTATE.
+ Admitted.
+ Hint Resolve transl_ijumptable_correct : htlproof.
+
+ Lemma transl_ireturn_correct:
+ forall (s : list RTL.stackframe) (f : RTL.function) (stk : Values.block)
+ (pc : positive) (rs : RTL.regset) (m : mem) (or : option Registers.reg)
+ (m' : mem),
+ (RTL.fn_code f) ! pc = Some (RTL.Ireturn or) ->
+ Mem.free m stk 0 (RTL.fn_stacksize f) = Some m' ->
+ forall R1 : HTL.state,
+ match_states (RTL.State s f (Values.Vptr stk Integers.Ptrofs.zero) pc rs m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\
+ match_states (RTL.Returnstate s (Registers.regmap_optget or Values.Vundef rs) m') R2.
+ Proof.
+ intros s f stk pc rs m or m' H H0 R1 MSTATE.
+ inv_state.
+
+ - econstructor. split.
eapply Smallstep.plus_two.
-
+
eapply HTL.step_module; eauto.
- apply assumption_32bit.
+ inv CONST; assumption.
+ inv CONST; assumption.
constructor.
econstructor; simpl; trivial.
econstructor; simpl; trivial.
@@ -2125,180 +2012,183 @@ Section CORRECTNESS.
constructor. constructor.
- unfold Verilog.merge_regs.
- unfold_merge. simpl.
- rewrite AssocMap.gso.
- rewrite AssocMap.gso.
- unfold state_st_wf in WF. eapply WF. reflexivity.
- apply st_greater_than_res. apply st_greater_than_res.
+ unfold state_st_wf in WF; big_tac; eauto.
+ destruct wf as [HCTRL HDATA]. apply HCTRL.
+ apply AssocMapExt.elements_iff. eexists.
+ match goal with H: control ! pc = Some _ |- _ => apply H end.
apply HTL.step_finish.
unfold Verilog.merge_regs.
unfold_merge; simpl.
rewrite AssocMap.gso.
- apply AssocMap.gss.
- apply finish_not_return.
+ apply AssocMap.gss. lia.
apply AssocMap.gss.
rewrite Events.E0_left. reflexivity.
- constructor.
- admit.
+ constructor; auto.
constructor.
+ (* FIXME: Duplication *)
- econstructor. split.
eapply Smallstep.plus_two.
eapply HTL.step_module; eauto.
- apply assumption_32bit.
+ inv CONST; assumption.
+ inv CONST; assumption.
constructor.
econstructor; simpl; trivial.
econstructor; simpl; trivial.
+ constructor. constructor. constructor.
+ constructor. constructor. constructor.
- constructor. constructor.
+ unfold state_st_wf in WF; big_tac; eauto.
- constructor.
- econstructor; simpl; trivial.
- apply Verilog.erun_Vvar. trivial.
- unfold Verilog.merge_regs.
- unfold_merge. simpl.
- rewrite AssocMap.gso.
- rewrite AssocMap.gso.
- unfold state_st_wf in WF. eapply WF. trivial.
- apply st_greater_than_res. apply st_greater_than_res. trivial.
+ destruct wf as [HCTRL HDATA]. apply HCTRL.
+ apply AssocMapExt.elements_iff. eexists.
+ match goal with H: control ! pc = Some _ |- _ => apply H end.
apply HTL.step_finish.
unfold Verilog.merge_regs.
unfold_merge.
rewrite AssocMap.gso.
- apply AssocMap.gss.
- apply finish_not_return.
+ apply AssocMap.gss. simpl; lia.
apply AssocMap.gss.
rewrite Events.E0_left. trivial.
- constructor.
- admit.
+ constructor; auto.
simpl. inversion MASSOC. subst.
unfold find_assocmap, AssocMapExt.get_default. rewrite AssocMap.gso.
apply H1. eapply RTL.max_reg_function_use. eauto. simpl; tauto.
- apply st_greater_than_res.
+ assert (HPle : Ple r (RTL.max_reg_function f)).
+ eapply RTL.max_reg_function_use. eassumption. simpl; auto.
+ apply ZExtra.Ple_not_eq. apply ZExtra.Ple_Plt_Succ. assumption.
- - inversion MSTATE; subst. inversion TF; subst.
- econstructor. split. apply Smallstep.plus_one.
- eapply HTL.step_call. simplify.
-
- apply match_state with (sp' := stk); eauto.
+ Unshelve.
+ all: constructor.
+ Qed.
+ Hint Resolve transl_ireturn_correct : htlproof.
+
+ Lemma transl_callstate_correct:
+ forall (s : list RTL.stackframe) (f : RTL.function) (args : list Values.val)
+ (m : mem) (m' : Mem.mem') (stk : Values.block),
+ Mem.alloc m 0 (RTL.fn_stacksize f) = (m', stk) ->
+ forall R1 : HTL.state,
+ match_states (RTL.Callstate s (AST.Internal f) args m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\
+ match_states
+ (RTL.State s f (Values.Vptr stk Integers.Ptrofs.zero) (RTL.fn_entrypoint f)
+ (RTL.init_regs args (RTL.fn_params f)) m') R2.
+ Proof.
+ intros s f args m m' stk H R1 MSTATE.
- apply regs_lessdef_add_greater.
- apply greater_than_max_func.
- apply init_reg_assoc_empty.
- unfold state_st_wf.
- intros. inv H3. apply AssocMap.gss. constructor.
-
- econstructor. simplify.
- repeat split. unfold HTL.empty_stack.
- simplify. apply AssocMap.gss.
- unfold arr_repeat. simplify.
- apply list_repeat_len.
- intros.
- destruct (Mem.load AST.Mint32 m' stk
- (Integers.Ptrofs.unsigned (Integers.Ptrofs.add
- Integers.Ptrofs.zero
- (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD.
- pose proof Mem.load_alloc_same as LOAD_ALLOC.
- pose proof H as ALLOC.
- eapply LOAD_ALLOC in ALLOC.
- 2: { exact LOAD. }
- rewrite ALLOC.
- repeat constructor.
- constructor.
+ inversion MSTATE; subst. inversion TF; subst.
+ econstructor. split. apply Smallstep.plus_one.
+ eapply HTL.step_call. crush.
- unfold reg_stack_based_pointers. intros.
- unfold RTL.init_regs; simplify.
- destruct (RTL.fn_params f);
- rewrite Registers.Regmap.gi; constructor.
-
- unfold arr_stack_based_pointers. intros.
- simplify.
- destruct (Mem.load AST.Mint32 m' stk
- (Integers.Ptrofs.unsigned (Integers.Ptrofs.add
- Integers.Ptrofs.zero
- (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD.
- pose proof Mem.load_alloc_same as LOAD_ALLOC.
- pose proof H as ALLOC.
- eapply LOAD_ALLOC in ALLOC.
- 2: { exact LOAD. }
- rewrite ALLOC.
- repeat constructor.
- constructor.
+ apply match_state with (sp' := stk); eauto.
- Transparent Mem.alloc. (* TODO: Since there are opaque there's probably a lemma. *)
- Transparent Mem.load.
- Transparent Mem.store.
- unfold stack_bounds.
- split.
+ all: big_tac.
- unfold Mem.alloc in H.
- invert H.
- simplify.
- unfold Mem.load.
- intros.
- match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence.
- invert v0. unfold Mem.range_perm in H3.
- unfold Mem.perm in H3. simplify.
- unfold Mem.perm_order' in H3.
- rewrite Integers.Ptrofs.add_zero_l in H3.
- rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia.
- exploit (H3 ptr). lia. intros.
- rewrite Maps.PMap.gss in H8.
- match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction.
- simplify.
- apply proj_sumbool_true in H10. lia.
-
- unfold Mem.alloc in H.
- invert H.
- simplify.
- unfold Mem.store.
- intros.
- match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence.
- invert v0. unfold Mem.range_perm in H3.
- unfold Mem.perm in H3. simplify.
- unfold Mem.perm_order' in H3.
- rewrite Integers.Ptrofs.add_zero_l in H3.
- rewrite Integers.Ptrofs.unsigned_repr in H3; simplify; try lia.
- exploit (H3 ptr). lia. intros.
- rewrite Maps.PMap.gss in H8.
- match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction.
- simplify.
- apply proj_sumbool_true in H10. lia.
- Opaque Mem.alloc.
- Opaque Mem.load.
- Opaque Mem.store.
-
- - invert MSTATE. invert MS.
- econstructor.
- split. apply Smallstep.plus_one.
- constructor.
+ apply regs_lessdef_add_greater. unfold Plt; lia.
+ apply regs_lessdef_add_greater. unfold Plt; lia.
+ apply regs_lessdef_add_greater. unfold Plt; lia.
+ apply init_reg_assoc_empty.
- constructor; auto.
- econstructor; auto.
- apply regs_lessdef_add_match; auto.
- apply regs_lessdef_add_greater. apply greater_than_max_func. auto.
+ constructor.
- unfold state_st_wf. intros. inv H. rewrite AssocMap.gso.
- rewrite AssocMap.gss. trivial. apply st_greater_than_res.
+ destruct (Mem.load AST.Mint32 m' stk
+ (Integers.Ptrofs.unsigned (Integers.Ptrofs.add
+ Integers.Ptrofs.zero
+ (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD.
+ pose proof Mem.load_alloc_same as LOAD_ALLOC.
+ pose proof H as ALLOC.
+ eapply LOAD_ALLOC in ALLOC.
+ 2: { exact LOAD. }
+ ptrofs. rewrite LOAD.
+ rewrite ALLOC.
+ repeat constructor.
+
+ ptrofs. rewrite LOAD.
+ repeat constructor.
+
+ unfold reg_stack_based_pointers. intros.
+ unfold RTL.init_regs; crush.
+ destruct (RTL.fn_params f);
+ rewrite Registers.Regmap.gi; constructor.
+
+ unfold arr_stack_based_pointers. intros.
+ crush.
+ destruct (Mem.load AST.Mint32 m' stk
+ (Integers.Ptrofs.unsigned (Integers.Ptrofs.add
+ Integers.Ptrofs.zero
+ (Integers.Ptrofs.repr (4 * ptr))))) eqn:LOAD.
+ pose proof Mem.load_alloc_same as LOAD_ALLOC.
+ pose proof H as ALLOC.
+ eapply LOAD_ALLOC in ALLOC.
+ 2: { exact LOAD. }
+ rewrite ALLOC.
+ repeat constructor.
+ constructor.
- admit.
+ Transparent Mem.alloc. (* TODO: Since there are opaque there's probably a lemma. *)
+ Transparent Mem.load.
+ Transparent Mem.store.
+ unfold stack_bounds.
+ split.
- Unshelve.
- exact (AssocMap.empty value).
- exact (AssocMap.empty value).
- (* exact (ZToValue 32 0). *)
- (* exact (AssocMap.empty value). *)
- (* exact (AssocMap.empty value). *)
- (* exact (AssocMap.empty value). *)
- (* exact (AssocMap.empty value). *)*)
- Admitted.
- Hint Resolve transl_step_correct : htlproof.
+ unfold Mem.alloc in H.
+ invert H.
+ crush.
+ unfold Mem.load.
+ intros.
+ match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence.
+ invert v0. unfold Mem.range_perm in H4.
+ unfold Mem.perm in H4. crush.
+ unfold Mem.perm_order' in H4.
+ small_tac.
+ exploit (H4 ptr). rewrite Integers.Ptrofs.unsigned_repr; small_tac. intros.
+ rewrite Maps.PMap.gss in H8.
+ match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction.
+ crush.
+ apply proj_sumbool_true in H10. lia.
+
+ unfold Mem.alloc in H.
+ invert H.
+ crush.
+ unfold Mem.store.
+ intros.
+ match goal with | |- context[if ?x then _ else _] => destruct x end; try congruence.
+ invert v0. unfold Mem.range_perm in H4.
+ unfold Mem.perm in H4. crush.
+ unfold Mem.perm_order' in H4.
+ small_tac.
+ exploit (H4 ptr). rewrite Integers.Ptrofs.unsigned_repr; small_tac. intros.
+ rewrite Maps.PMap.gss in H8.
+ match goal with | H8 : context[if ?x then _ else _] |- _ => destruct x eqn:EQ end; try contradiction.
+ crush.
+ apply proj_sumbool_true in H10. lia.
+ constructor. simplify. rewrite AssocMap.gss.
+ simplify. rewrite AssocMap.gso. apply AssocMap.gss. simplify. lia.
+ Opaque Mem.alloc.
+ Opaque Mem.load.
+ Opaque Mem.store.
+ Qed.
+ Hint Resolve transl_callstate_correct : htlproof.
+
+ Lemma transl_returnstate_correct:
+ forall (res0 : Registers.reg) (f : RTL.function) (sp : Values.val) (pc : RTL.node)
+ (rs : RTL.regset) (s : list RTL.stackframe) (vres : Values.val) (m : mem)
+ (R1 : HTL.state),
+ match_states (RTL.Returnstate (RTL.Stackframe res0 f sp pc rs :: s) vres m) R1 ->
+ exists R2 : HTL.state,
+ Smallstep.plus HTL.step tge R1 Events.E0 R2 /\
+ match_states (RTL.State s f sp pc (Registers.Regmap.set res0 vres rs) m) R2.
+ Proof.
+ intros res0 f sp pc rs s vres m R1 MSTATE.
+ inversion MSTATE. inversion MF.
+ Qed.
+ Hint Resolve transl_returnstate_correct : htlproof.
Lemma option_inv :
forall A x y,
@@ -2322,8 +2212,6 @@ Section CORRECTNESS.
trivial. symmetry; eapply Linking.match_program_main; eauto.
Qed.
-
- (* Had to admit proof because currently there is no way to force main to be Internal. *)
Lemma transl_initial_states :
forall s1 : Smallstep.state (RTL.semantics prog),
Smallstep.initial_state (RTL.semantics prog) s1 ->
@@ -2370,18 +2258,26 @@ Section CORRECTNESS.
Smallstep.final_state (RTL.semantics prog) s1 r ->
Smallstep.final_state (HTL.semantics tprog) s2 r.
Proof.
- intros. inv H0. inv H. inv H4. inv MS. constructor. trivial.
+ intros. inv H0. inv H. inv H4. invert MF. constructor. reflexivity.
Qed.
Hint Resolve transl_final_states : htlproof.
-Theorem transf_program_correct:
- Smallstep.forward_simulation (RTL.semantics prog) (HTL.semantics tprog).
-Proof.
- eapply Smallstep.forward_simulation_plus.
- apply senv_preserved.
- eexact transl_initial_states.
- eexact transl_final_states.
- exact transl_step_correct.
-Qed.
+ 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, Smallstep.plus HTL.step tge R1 t R2 /\ match_states S2 R2.
+ Proof.
+ induction 1; eauto with htlproof; (intros; inv_state).
+ Qed.
+ Hint Resolve transl_step_correct : htlproof.
+
+ Theorem transf_program_correct:
+ Smallstep.forward_simulation (RTL.semantics prog) (HTL.semantics tprog).
+ Proof.
+ eapply Smallstep.forward_simulation_plus; eauto with htlproof.
+ apply senv_preserved.
+ Qed.
End CORRECTNESS.
diff --git a/src/translation/HTLgenspec.v b/src/translation/HTLgenspec.v
index a78256b..1b04b1f 100644
--- a/src/translation/HTLgenspec.v
+++ b/src/translation/HTLgenspec.v
@@ -17,8 +17,9 @@
*)
From compcert Require RTL Op Maps Errors.
-From compcert Require Import Maps.
-From coqup Require Import Coquplib Verilog Value HTL HTLgen AssocMap.
+From compcert Require Import Maps Integers.
+From coqup Require Import Coquplib Verilog ValueInt HTL HTLgen AssocMap.
+Require Import Lia.
Hint Resolve Maps.PTree.elements_keys_norepet : htlspec.
Hint Resolve Maps.PTree.elements_correct : htlspec.
@@ -116,31 +117,40 @@ translations for each of the elements *)
Inductive tr_instr (fin rtrn st stk : reg) : RTL.instruction -> stmnt -> stmnt -> Prop :=
| tr_instr_Inop :
forall n,
+ Z.pos n <= Int.max_unsigned ->
tr_instr fin rtrn st stk (RTL.Inop n) Vskip (state_goto st n)
| tr_instr_Iop :
forall n op args dst s s' e i,
+ Z.pos n <= Int.max_unsigned ->
translate_instr op args s = OK e s' i ->
tr_instr fin rtrn st stk (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,
+ Z.pos n1 <= Int.max_unsigned ->
+ Z.pos n2 <= Int.max_unsigned ->
translate_condition cond args s = OK c s' i ->
tr_instr fin rtrn st stk (RTL.Icond cond args n1 n2) Vskip (state_cond st c n1 n2)
| tr_instr_Ireturn_None :
- tr_instr fin rtrn st stk (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 stk (RTL.Ireturn None) (Vseq (block fin (Vlit (ZToValue 1%Z)))
+ (block rtrn (Vlit (ZToValue 0%Z)))) Vskip
| tr_instr_Ireturn_Some :
forall r,
tr_instr fin rtrn st stk (RTL.Ireturn (Some r))
- (Vseq (block fin (Vlit (ZToValue 1%nat 1%Z))) (block rtrn (Vvar r))) Vskip
+ (Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vvar r))) Vskip
| tr_instr_Iload :
- forall mem addr args s s' i c dst n,
- translate_arr_access mem addr args stk s = OK c s' i ->
- tr_instr fin rtrn st stk (RTL.Iload mem addr args dst n) (nonblock dst c) (state_goto st n)
+ forall chunk addr args s s' i e dst n,
+ Z.pos n <= Int.max_unsigned ->
+ chunk = AST.Mint32 ->
+ translate_arr_addressing addr args s = OK e s' i ->
+ tr_instr fin rtrn st stk (RTL.Iload chunk addr args dst n)
+ (create_single_cycle_load stk e dst) (state_goto st n)
| tr_instr_Istore :
- forall mem addr args s s' i c src n,
- translate_arr_access mem addr args stk s = OK c s' i ->
- tr_instr fin rtrn st stk (RTL.Istore mem addr args src n) (Vnonblock c (Vvar src))
- (state_goto st n)
+ forall chunk addr args s s' i e src n,
+ Z.pos n <= Int.max_unsigned ->
+ chunk = AST.Mint32 ->
+ translate_arr_addressing addr args s = OK e s' i ->
+ tr_instr fin rtrn st stk (RTL.Istore chunk addr args src n)
+ (create_single_cycle_store stk e src) (state_goto st n)
| tr_instr_Ijumptable :
forall cexpr tbl r,
cexpr = tbl_to_case_expr st tbl ->
@@ -160,17 +170,24 @@ Hint Constructors tr_code : htlspec.
Inductive tr_module (f : RTL.function) : module -> Prop :=
tr_module_intro :
- forall data control fin rtrn st stk stk_len m start rst clk scldecls arrdecls,
- (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i ->
- tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) ->
- stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) ->
- Z.modulo (f.(RTL.fn_stacksize)) 4 = 0 ->
- 0 <= f.(RTL.fn_stacksize) < Integers.Ptrofs.modulus ->
+ forall data control fin rtrn st stk stk_len m start rst clk scldecls arrdecls wf,
m = (mkmodule f.(RTL.fn_params)
data
control
f.(RTL.fn_entrypoint)
- st stk stk_len fin rtrn start rst clk scldecls arrdecls) ->
+ st stk stk_len fin rtrn start rst clk scldecls arrdecls wf) ->
+ (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i ->
+ tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) ->
+ stk_len = Z.to_nat f.(RTL.fn_stacksize) ->
+ Z.modulo (f.(RTL.fn_stacksize)) 4 = 0 ->
+ 0 <= f.(RTL.fn_stacksize) < Integers.Ptrofs.modulus ->
+ st = ((RTL.max_reg_function f) + 1)%positive ->
+ fin = ((RTL.max_reg_function f) + 2)%positive ->
+ rtrn = ((RTL.max_reg_function f) + 3)%positive ->
+ stk = ((RTL.max_reg_function f) + 4)%positive ->
+ start = ((RTL.max_reg_function f) + 5)%positive ->
+ rst = ((RTL.max_reg_function f) + 6)%positive ->
+ clk = ((RTL.max_reg_function f) + 7)%positive ->
tr_module f m.
Hint Constructors tr_module : htlspec.
@@ -202,6 +219,13 @@ Lemma declare_reg_controllogic_trans :
Proof. intros. monadInv H. trivial. Qed.
Hint Resolve create_reg_controllogic_trans : htlspec.
+Lemma declare_reg_freshreg_trans :
+ forall sz s s' x i iop r,
+ declare_reg iop r sz s = OK x s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof. inversion 1; auto. Qed.
+Hint Resolve declare_reg_freshreg_trans : htlspec.
+
Lemma create_arr_datapath_trans :
forall sz ln s s' x i iop,
create_arr iop sz ln s = OK x s' i ->
@@ -268,6 +292,16 @@ Proof.
- apply H in EQ. rewrite EQ. eauto.
Qed.
+Lemma collect_freshreg_trans :
+ forall A f l cs cs' ci,
+ (forall s s' x i y, f y s = OK x s' i -> s.(st_freshreg) = s'.(st_freshreg)) ->
+ @HTLMonadExtra.collectlist A f l cs = OK tt cs' ci -> cs.(st_freshreg) = cs'.(st_freshreg).
+Proof.
+ induction l; intros; monadInv H0.
+ - trivial.
+ - apply H in EQ. rewrite EQ. eauto.
+Qed.
+
Lemma collect_declare_controllogic_trans :
forall io n l s s' i,
HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i ->
@@ -286,6 +320,135 @@ Proof.
intros. eapply declare_reg_datapath_trans. simpl in H0. eassumption.
Qed.
+Lemma collect_declare_freshreg_trans :
+ forall io n l s s' i,
+ HTLMonadExtra.collectlist (fun r : reg => declare_reg io r n) l s = OK tt s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ intros. eapply collect_freshreg_trans; try eassumption.
+ inversion 1. auto.
+Qed.
+
+Ltac unfold_match H :=
+ match type of H with
+ | context[match ?g with _ => _ end] => destruct g eqn:?; try discriminate
+ end.
+
+Lemma translate_eff_addressing_freshreg_trans :
+ forall op args s r s' i,
+ translate_eff_addressing op args s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto.
+Qed.
+Hint Resolve translate_eff_addressing_freshreg_trans : htlspec.
+
+Lemma translate_arr_addressing_freshreg_trans :
+ forall op args s r s' i,
+ translate_arr_addressing op args s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto.
+Qed.
+Hint Resolve translate_eff_addressing_freshreg_trans : htlspec.
+
+Lemma translate_comparison_freshreg_trans :
+ forall op args s r s' i,
+ translate_comparison op args s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto.
+Qed.
+Hint Resolve translate_comparison_freshreg_trans : htlspec.
+
+Lemma translate_comparison_imm_freshreg_trans :
+ forall op args s r s' i n,
+ translate_comparison_imm op args n s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ destruct op; intros; simpl in *; repeat (unfold_match H); inv H; auto.
+Qed.
+Hint Resolve translate_comparison_imm_freshreg_trans : htlspec.
+
+Lemma translate_condition_freshreg_trans :
+ forall op args s r s' i,
+ translate_condition op args s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ destruct op; intros; simpl in *; repeat (unfold_match H); inv H; eauto with htlspec.
+Qed.
+Hint Resolve translate_condition_freshreg_trans : htlspec.
+
+Lemma translate_instr_freshreg_trans :
+ forall op args s r s' i,
+ translate_instr op args s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ destruct op; intros; simpl in *; repeat (unfold_match H); inv H; eauto with htlspec.
+ monadInv H1. eauto with htlspec.
+Qed.
+Hint Resolve translate_instr_freshreg_trans : htlspec.
+
+Lemma add_instr_freshreg_trans :
+ forall n n' st s r s' i,
+ add_instr n n' st s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof. intros. unfold add_instr in H. repeat (unfold_match H). inv H. auto. Qed.
+Hint Resolve add_instr_freshreg_trans : htlspec.
+
+Lemma add_branch_instr_freshreg_trans :
+ forall n n0 n1 e s r s' i,
+ add_branch_instr e n n0 n1 s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof. intros. unfold add_branch_instr in H. repeat (unfold_match H). inv H. auto. Qed.
+Hint Resolve add_branch_instr_freshreg_trans : htlspec.
+
+Lemma add_node_skip_freshreg_trans :
+ forall n1 n2 s r s' i,
+ add_node_skip n1 n2 s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof. intros. unfold add_node_skip in H. repeat (unfold_match H). inv H. auto. Qed.
+Hint Resolve add_node_skip_freshreg_trans : htlspec.
+
+Lemma add_instr_skip_freshreg_trans :
+ forall n1 n2 s r s' i,
+ add_instr_skip n1 n2 s = OK r s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof. intros. unfold add_instr_skip in H. repeat (unfold_match H). inv H. auto. Qed.
+Hint Resolve add_instr_skip_freshreg_trans : htlspec.
+
+Lemma transf_instr_freshreg_trans :
+ forall fin ret st instr s v s' i,
+ transf_instr fin ret st instr s = OK v s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ intros. destruct instr eqn:?. subst. unfold transf_instr in H.
+ destruct i0; try (monadInv H); try (unfold_match H); eauto with htlspec.
+ - monadInv H. apply add_instr_freshreg_trans in EQ2. apply translate_instr_freshreg_trans in EQ.
+ apply declare_reg_freshreg_trans in EQ1. congruence.
+ - destruct (Z.pos n0 <=? Int.max_unsigned); try discriminate.
+ monadInv H. apply add_instr_freshreg_trans in EQ2.
+ apply translate_arr_addressing_freshreg_trans in EQ.
+ apply declare_reg_freshreg_trans in EQ1. congruence.
+ - destruct (Z.pos n0 <=? Int.max_unsigned); try discriminate.
+ monadInv H. apply add_instr_freshreg_trans in EQ0.
+ apply translate_arr_addressing_freshreg_trans in EQ. congruence.
+ - monadInv H. apply translate_condition_freshreg_trans in EQ.
+ apply add_branch_instr_freshreg_trans in EQ0.
+ congruence.
+ - inv EQ. apply add_node_skip_freshreg_trans in EQ0. congruence.
+Qed.
+Hint Resolve transf_instr_freshreg_trans : htlspec.
+
+Lemma collect_trans_instr_freshreg_trans :
+ forall fin ret st l s s' i,
+ HTLMonadExtra.collectlist (transf_instr fin ret st) l s = OK tt s' i ->
+ s.(st_freshreg) = s'.(st_freshreg).
+Proof.
+ intros. eapply collect_freshreg_trans; try eassumption.
+ eauto with htlspec.
+Qed.
+
Ltac rewrite_states :=
match goal with
| [ H: ?x ?s = ?x ?s' |- _ ] =>
@@ -294,20 +457,18 @@ Ltac rewrite_states :=
remember (?x ?s) as c1; remember (?x ?s') as c2; try subst
end.
-Ltac unfold_match H :=
- match type of H with
- | context[match ?g with _ => _ end] => destruct g eqn:?; 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
| ?f _ _ _ _ _ = OK _ _ _ => unfold f in H
+ | ?f _ _ _ _ _ _ = OK _ _ _ => unfold f in H
end; repeat unfold_match H; inversion H.
Ltac inv_add_instr :=
- lazymatch goal with
+ match goal with
+ | H: (if ?c then _ else _) _ = OK _ _ _ |- _ => destruct c eqn:EQN; try discriminate; inv_add_instr
| H: context[add_instr_skip _ _ _] |- _ =>
inv_add_instr' H
| H: context[add_instr_skip _ _] |- _ =>
@@ -343,28 +504,34 @@ Proof.
destruct (peq pc pc1).
- subst.
destruct instr1 eqn:?; try discriminate;
- try destruct_optional; inv_add_instr; econstructor; try assumption.
+ try destruct_optional; try (destruct m; try discriminate);
+ inv_add_instr; econstructor; try assumption.
+ destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence.
+ destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence.
+ inversion H2. inversion H9. rewrite H. apply tr_instr_Inop.
+ apply Z.leb_le. assumption.
eapply in_map with (f := fst) in H9. contradiction.
+ destruct o with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence.
+ destruct o0 with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence.
+ inversion H2. inversion H14. unfold nonblock. replace (st_st s4) with (st_st s2) by congruence.
- econstructor. apply EQ1. eapply in_map with (f := fst) in H14. contradiction.
+ econstructor. apply Z.leb_le; assumption.
+ apply EQ1. eapply in_map with (f := fst) in H14. contradiction.
+ destruct o with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence.
+ destruct o0 with pc1; destruct H16; simpl in *; rewrite AssocMap.gss in H14; eauto; congruence.
+ inversion H2. inversion H14. rewrite <- e2. replace (st_st s2) with (st_st s0) by congruence.
- econstructor. apply EQ1. eapply in_map with (f := fst) in H14. contradiction.
+ econstructor. apply Z.leb_le; assumption.
+ reflexivity.
+ apply EQ1. eapply in_map with (f := fst) in H14. contradiction.
+ destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence.
+ destruct o0 with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence.
+ destruct H2.
* inversion H2.
replace (st_st s2) with (st_st s0) by congruence.
- eauto with htlspec.
+ econstructor. apply Z.leb_le; assumption.
+ eauto with htlspec. eassumption.
* apply in_map with (f := fst) in H2. contradiction.
+ destruct o with pc1; destruct H11; simpl in *; rewrite AssocMap.gss in H9; eauto; congruence.
@@ -372,6 +539,7 @@ Proof.
+ destruct H2.
* inversion H2.
replace (st_st s2) with (st_st s0) by congruence.
+ econstructor; try (apply Z.leb_le; apply andb_prop in EQN; apply EQN).
eauto with htlspec.
* apply in_map with (f := fst) in H2. contradiction.
@@ -405,9 +573,17 @@ Qed.
Hint Resolve iter_expand_instr_spec : htlspec.
Lemma create_arr_inv : forall w x y z a b c d,
- create_arr w x y z = OK (a, b) c d -> y = b.
+ create_arr w x y z = OK (a, b) c d ->
+ y = b /\ a = z.(st_freshreg) /\ c.(st_freshreg) = Pos.succ (z.(st_freshreg)).
+Proof.
+ inversion 1; split; auto.
+Qed.
+
+Lemma create_reg_inv : forall a b s r s' i,
+ create_reg a b s = OK r s' i ->
+ r = s.(st_freshreg) /\ s'.(st_freshreg) = Pos.succ (s.(st_freshreg)).
Proof.
- inversion 1. reflexivity.
+ inversion 1; auto.
Qed.
Theorem transl_module_correct :
@@ -426,24 +602,32 @@ Proof.
destruct (0 <=? RTL.fn_stacksize f) eqn:STACK_BOUND_LOW;
destruct (RTL.fn_stacksize f <? Integers.Ptrofs.modulus) eqn:STACK_BOUND_HIGH;
destruct (RTL.fn_stacksize f mod 4 =? 0) eqn:STACK_ALIGN;
- simplify;
+ crush;
monadInv Heqr.
- (* TODO: We should be able to fold this into the automation. *)
- pose proof (create_arr_inv _ _ _ _ _ _ _ _ EQ0) as STK_LEN.
- rewrite <- STK_LEN.
+ repeat unfold_match EQ9. monadInv EQ9.
- econstructor; simpl; auto.
- intros.
+ (* TODO: We should be able to fold this into the automation. *)
+ pose proof (create_arr_inv _ _ _ _ _ _ _ _ EQ0) as STK_LEN. inv STK_LEN. inv H5.
+ pose proof (create_reg_inv _ _ _ _ _ _ EQ) as FIN_VAL. inv FIN_VAL.
+ pose proof (create_reg_inv _ _ _ _ _ _ EQ1) as RET_VAL. inv RET_VAL.
+ destruct x3. destruct x4.
+ pose proof (collect_trans_instr_freshreg_trans _ _ _ _ _ _ _ EQ2) as TR_INSTR.
+ pose proof (collect_declare_freshreg_trans _ _ _ _ _ _ EQ3) as TR_DEC.
+ pose proof (create_reg_inv _ _ _ _ _ _ EQ4) as START_VAL. inv START_VAL.
+ pose proof (create_reg_inv _ _ _ _ _ _ EQ5) as RESET_VAL. inv RESET_VAL.
+ pose proof (create_reg_inv _ _ _ _ _ _ EQ6) as CLK_VAL. inv CLK_VAL.
+ rewrite H9 in *. rewrite H8 in *. replace (st_freshreg s4) with (st_freshreg s2) in * by congruence.
+ rewrite H6 in *. rewrite H7 in *. rewrite H5 in *. simpl in *.
inv_incr.
+ econstructor; simpl; auto; try lia.
+ intros.
assert (EQ3D := EQ3).
- destruct x4.
apply collect_declare_datapath_trans in EQ3.
apply collect_declare_controllogic_trans in EQ3D.
- assert (STC: st_controllogic s10 = st_controllogic s3) by congruence.
- assert (STD: st_datapath s10 = st_datapath s3) by congruence.
- assert (STST: st_st s10 = st_st s3) by congruence.
- rewrite STC. rewrite STD. rewrite STST.
+ replace (st_controllogic s10) with (st_controllogic s3) by congruence.
+ replace (st_datapath s10) with (st_datapath s3) by congruence.
+ replace (st_st s10) with (st_st s3) by congruence.
eapply iter_expand_instr_spec; eauto with htlspec.
apply PTree.elements_complete.
Qed.
diff --git a/src/translation/Veriloggen.v b/src/translation/Veriloggen.v
index b550ff9..f5d5fa7 100644
--- a/src/translation/Veriloggen.v
+++ b/src/translation/Veriloggen.v
@@ -19,35 +19,33 @@
From compcert Require Import Maps.
From compcert Require Errors.
From compcert Require Import AST.
-From coqup Require Import Verilog HTL Coquplib AssocMap Value.
+From coqup Require Import Verilog HTL Coquplib AssocMap ValueInt.
-Fixpoint transl_list (st : list (node * Verilog.stmnt)) {struct st} : list (expr * Verilog.stmnt) :=
- match st with
- | (n, stmt) :: ls => (Vlit (posToValue 32 n), stmt) :: transl_list ls
- | nil => nil
- end.
+Definition transl_list_fun (a : node * Verilog.stmnt) :=
+ let (n, stmnt) := a in
+ (Vlit (posToValue n), stmnt).
-Fixpoint scl_to_Vdecl (scldecl : list (reg * (option io * scl_decl))) {struct scldecl} : list module_item :=
- match scldecl with
- | (r, (io, VScalar sz))::scldecl' => Vdeclaration (Vdecl io r sz) :: scl_to_Vdecl scldecl'
- | nil => nil
- end.
+Definition transl_list st := map transl_list_fun st.
-Fixpoint arr_to_Vdeclarr (arrdecl : list (reg * (option io * arr_decl))) {struct arrdecl} : list module_item :=
- match arrdecl with
- | (r, (io, VArray sz l))::arrdecl' => Vdeclaration (Vdeclarr io r sz l) :: arr_to_Vdeclarr arrdecl'
- | nil => nil
- end.
+Definition scl_to_Vdecl_fun (a : reg * (option io * scl_decl)) :=
+ match a with (r, (io, VScalar sz)) => (Vdecl io r sz) end.
+
+Definition scl_to_Vdecl scldecl := map scl_to_Vdecl_fun scldecl.
+
+Definition arr_to_Vdeclarr_fun (a : reg * (option io * arr_decl)) :=
+ match a with (r, (io, VArray sz l)) => (Vdeclarr io r sz l) end.
+
+Definition arr_to_Vdeclarr arrdecl := map arr_to_Vdeclarr_fun arrdecl.
Definition transl_module (m : HTL.module) : Verilog.module :=
let case_el_ctrl := transl_list (PTree.elements m.(mod_controllogic)) in
let case_el_data := transl_list (PTree.elements m.(mod_datapath)) in
let body :=
- Valways (Vposedge m.(mod_clk)) (Vcond (Vbinop Veq (Vvar m.(mod_reset)) (ZToValue 1 1))
- (Vnonblock (Vvar m.(mod_st)) (posToValue 32 m.(mod_entrypoint)))
+ Valways (Vposedge m.(mod_clk)) (Vcond (Vbinop Veq (Vvar m.(mod_reset)) (Vlit (ZToValue 1)))
+ (Vnonblock (Vvar m.(mod_st)) (Vlit (posToValue m.(mod_entrypoint))))
(Vcase (Vvar m.(mod_st)) case_el_ctrl (Some Vskip)))
:: Valways (Vposedge m.(mod_clk)) (Vcase (Vvar m.(mod_st)) case_el_data (Some Vskip))
- :: (arr_to_Vdeclarr (AssocMap.elements m.(mod_arrdecls))
+ :: List.map Vdeclaration (arr_to_Vdeclarr (AssocMap.elements m.(mod_arrdecls))
++ scl_to_Vdecl (AssocMap.elements m.(mod_scldecls))) in
Verilog.mkmodule m.(mod_start)
m.(mod_reset)
diff --git a/src/translation/Veriloggenproof.v b/src/translation/Veriloggenproof.v
index ca4ecab..5b467a7 100644
--- a/src/translation/Veriloggenproof.v
+++ b/src/translation/Veriloggenproof.v
@@ -16,9 +16,12 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-From compcert Require Import Smallstep Linking.
+From compcert Require Import Smallstep Linking Integers Globalenvs.
From coqup Require HTL.
-From coqup Require Import Coquplib Veriloggen Verilog.
+From coqup Require Import Coquplib Veriloggen Verilog ValueInt AssocMap.
+Require Import Lia.
+
+Local Open Scope assocmap.
Definition match_prog (prog : HTL.program) (tprog : Verilog.program) :=
match_program (fun cu f tf => tf = transl_fundef f) eq prog tprog.
@@ -52,6 +55,176 @@ Inductive match_states : HTL.state -> state -> Prop :=
forall m,
match_states (HTL.Callstate nil m nil) (Callstate nil (transl_module m) nil).
+Lemma Vlit_inj :
+ forall a b, Vlit a = Vlit b -> a = b.
+Proof. inversion 1. trivial. Qed.
+
+Lemma posToValue_inj :
+ forall a b,
+ 0 <= Z.pos a <= Int.max_unsigned ->
+ 0 <= Z.pos b <= Int.max_unsigned ->
+ posToValue a = posToValue b ->
+ a = b.
+Proof.
+ intros. rewrite <- Pos2Z.id at 1. rewrite <- Pos2Z.id.
+ rewrite <- Int.unsigned_repr at 1; try assumption.
+ symmetry.
+ rewrite <- Int.unsigned_repr at 1; try assumption.
+ unfold posToValue in *. rewrite H1; auto.
+Qed.
+
+Lemma valueToPos_inj :
+ forall a b,
+ 0 < Int.unsigned a ->
+ 0 < Int.unsigned b ->
+ valueToPos a = valueToPos b ->
+ a = b.
+Proof.
+ intros. unfold valueToPos in *.
+ rewrite <- Int.repr_unsigned at 1.
+ rewrite <- Int.repr_unsigned.
+ apply Pos2Z.inj_iff in H1.
+ rewrite Z2Pos.id in H1; auto.
+ rewrite Z2Pos.id in H1; auto.
+ rewrite H1. auto.
+Qed.
+
+Lemma unsigned_posToValue_le :
+ forall p,
+ Z.pos p <= Int.max_unsigned ->
+ 0 < Int.unsigned (posToValue p).
+Proof.
+ intros. unfold posToValue. rewrite Int.unsigned_repr; lia.
+Qed.
+
+Lemma transl_list_fun_fst :
+ forall p1 p2 a b,
+ 0 <= Z.pos p1 <= Int.max_unsigned ->
+ 0 <= Z.pos p2 <= Int.max_unsigned ->
+ transl_list_fun (p1, a) = transl_list_fun (p2, b) ->
+ (p1, a) = (p2, b).
+Proof.
+ intros. unfold transl_list_fun in H1. apply pair_equal_spec in H1.
+ destruct H1. rewrite H2. apply Vlit_inj in H1.
+ apply posToValue_inj in H1; try assumption.
+ rewrite H1; auto.
+Qed.
+
+Lemma Zle_relax :
+ forall p q r,
+ p < q <= r ->
+ p <= q <= r.
+Proof. lia. Qed.
+Hint Resolve Zle_relax : verilogproof.
+
+Lemma transl_in :
+ forall l p,
+ 0 <= Z.pos p <= Int.max_unsigned ->
+ (forall p0, In p0 (List.map fst l) -> 0 < Z.pos p0 <= Int.max_unsigned) ->
+ In (Vlit (posToValue p)) (map fst (map transl_list_fun l)) ->
+ In p (map fst l).
+Proof.
+ induction l.
+ - simplify. auto.
+ - intros. destruct a. simplify. destruct (peq p0 p); auto.
+ right. inv H1. apply Vlit_inj in H. apply posToValue_inj in H; auto. contradiction.
+ auto with verilogproof.
+ apply IHl; auto.
+Qed.
+
+Lemma transl_notin :
+ forall l p,
+ 0 <= Z.pos p <= Int.max_unsigned ->
+ (forall p0, In p0 (List.map fst l) -> 0 < Z.pos p0 <= Int.max_unsigned) ->
+ ~ In p (List.map fst l) -> ~ In (Vlit (posToValue p)) (List.map fst (transl_list l)).
+Proof.
+ induction l; auto. intros. destruct a. unfold not in *. intros.
+ simplify.
+ destruct (peq p0 p). apply H1. auto. apply H1.
+ unfold transl_list in *. inv H2. apply Vlit_inj in H. apply posToValue_inj in H.
+ contradiction.
+ auto with verilogproof. auto.
+ right. apply transl_in; auto.
+Qed.
+
+Lemma transl_norepet :
+ forall l,
+ (forall p0, In p0 (List.map fst l) -> 0 < Z.pos p0 <= Int.max_unsigned) ->
+ list_norepet (List.map fst l) -> list_norepet (List.map fst (transl_list l)).
+Proof.
+ induction l.
+ - intros. simpl. apply list_norepet_nil.
+ - destruct a. intros. simpl. apply list_norepet_cons.
+ inv H0. apply transl_notin. apply Zle_relax. apply H. simplify; auto.
+ intros. apply H. destruct (peq p0 p); subst; simplify; auto.
+ assumption. apply IHl. intros. apply H. destruct (peq p0 p); subst; simplify; auto.
+ simplify. inv H0. assumption.
+Qed.
+
+Lemma transl_list_correct :
+ forall l v ev f asr asa asrn asan asr' asa' asrn' asan',
+ (forall p0, In p0 (List.map fst l) -> 0 < Z.pos p0 <= Int.max_unsigned) ->
+ list_norepet (List.map fst l) ->
+ asr!ev = Some v ->
+ (forall p s,
+ In (p, s) l ->
+ v = posToValue p ->
+ stmnt_runp f
+ {| assoc_blocking := asr; assoc_nonblocking := asrn |}
+ {| assoc_blocking := asa; assoc_nonblocking := asan |}
+ s
+ {| assoc_blocking := asr'; assoc_nonblocking := asrn' |}
+ {| assoc_blocking := asa'; assoc_nonblocking := asan' |} ->
+ stmnt_runp f
+ {| assoc_blocking := asr; assoc_nonblocking := asrn |}
+ {| assoc_blocking := asa; assoc_nonblocking := asan |}
+ (Vcase (Vvar ev) (transl_list l) (Some Vskip))
+ {| assoc_blocking := asr'; assoc_nonblocking := asrn' |}
+ {| assoc_blocking := asa'; assoc_nonblocking := asan' |}).
+Proof.
+ induction l as [| a l IHl].
+ - contradiction.
+ - intros v ev f asr asa asrn asan asr' asa' asrn' asan' BOUND NOREP ASSOC p s IN EQ SRUN.
+ destruct a as [p' s']. simplify.
+ destruct (peq p p'); subst. eapply stmnt_runp_Vcase_match.
+ constructor. simplify. unfold AssocMap.find_assocmap, AssocMapExt.get_default.
+ rewrite ASSOC. trivial. constructor. auto.
+ inversion IN as [INV | INV]. inversion INV as [INV2]; subst. assumption.
+ inv NOREP. eapply in_map with (f := fst) in INV. contradiction.
+
+ eapply stmnt_runp_Vcase_nomatch. constructor. simplify.
+ unfold AssocMap.find_assocmap, AssocMapExt.get_default. rewrite ASSOC.
+ trivial. constructor. unfold not. intros. apply n. apply posToValue_inj.
+ apply Zle_relax. apply BOUND. right. inv IN. inv H0; contradiction.
+ eapply in_map with (f := fst) in H0. auto.
+ apply Zle_relax. apply BOUND; auto. auto.
+
+ eapply IHl. auto. inv NOREP. auto. eassumption. inv IN. inv H. contradiction. apply H.
+ trivial. assumption.
+Qed.
+Hint Resolve transl_list_correct : verilogproof.
+
+Lemma pc_wf :
+ forall A m p v,
+ (forall p0, In p0 (List.map fst (@AssocMap.elements A m)) -> 0 < Z.pos p0 <= Int.max_unsigned) ->
+ m!p = Some v ->
+ 0 <= Z.pos p <= Int.max_unsigned.
+Proof.
+ intros A m p v LT S. apply Zle_relax. apply LT.
+ apply AssocMap.elements_correct in S. remember (p, v) as x.
+ exploit in_map. apply S. instantiate (1 := fst). subst. simplify. auto.
+Qed.
+
+Lemma mis_stepp_decl :
+ forall l asr asa f,
+ mis_stepp f asr asa (map Vdeclaration l) asr asa.
+Proof.
+ induction l.
+ - intros. constructor.
+ - intros. simplify. econstructor. constructor. auto.
+Qed.
+Hint Resolve mis_stepp_decl : verilogproof.
+
Section CORRECTNESS.
Variable prog: HTL.program.
@@ -62,6 +235,40 @@ Section CORRECTNESS.
Let ge : HTL.genv := Globalenvs.Genv.globalenv prog.
Let tge : genv := Globalenvs.Genv.globalenv tprog.
+ Lemma symbols_preserved:
+ forall (s: AST.ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
+ Proof. intros. eapply (Genv.find_symbol_match TRANSL). Qed.
+ Hint Resolve symbols_preserved : verilogproof.
+
+ Lemma function_ptr_translated:
+ forall (b: Values.block) (f: HTL.fundef),
+ Genv.find_funct_ptr ge b = Some f ->
+ exists tf,
+ Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = tf.
+ Proof.
+ intros. exploit (Genv.find_funct_ptr_match TRANSL); eauto.
+ intros (cu & tf & P & Q & R); exists tf; auto.
+ Qed.
+ Hint Resolve function_ptr_translated : verilogproof.
+
+ Lemma functions_translated:
+ forall (v: Values.val) (f: HTL.fundef),
+ Genv.find_funct ge v = Some f ->
+ exists tf,
+ Genv.find_funct tge v = Some tf /\ transl_fundef f = tf.
+ Proof.
+ intros. exploit (Genv.find_funct_match TRANSL); eauto.
+ intros (cu & tf & P & Q & R); exists tf; auto.
+ Qed.
+ Hint Resolve functions_translated : verilogproof.
+
+ Lemma senv_preserved:
+ Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge).
+ Proof.
+ intros. eapply (Genv.senv_match TRANSL).
+ Qed.
+ Hint Resolve senv_preserved : verilogproof.
+
Theorem transl_step_correct :
forall (S1 : HTL.state) t S2,
HTL.step ge S1 t S2 ->
@@ -69,15 +276,93 @@ Section CORRECTNESS.
match_states S1 R1 ->
exists R2, Smallstep.plus step tge R1 t R2 /\ match_states S2 R2.
Proof.
-(* induction 1; intros R1 MSTATE; inv MSTATE; econstructor; split.
- - apply Smallstep.plus_one. econstructor. eassumption. trivial.
- * econstructor. econstructor.*)
- Admitted.
+ induction 1; intros R1 MSTATE; inv MSTATE.
+ - econstructor; split. apply Smallstep.plus_one. econstructor.
+ assumption. assumption. eassumption. apply valueToPos_posToValue.
+ split. lia.
+ eapply pc_wf. intros. pose proof (HTL.mod_wf m) as HP. destruct HP as [HP _].
+ split. lia. apply HP. eassumption. eassumption.
+ econstructor. econstructor. eapply stmnt_runp_Vcond_false. econstructor. econstructor.
+ simpl. unfold find_assocmap. unfold AssocMapExt.get_default.
+ rewrite H. trivial.
+
+ econstructor. simpl. auto. auto.
+
+ eapply transl_list_correct.
+ intros. split. lia. pose proof (HTL.mod_wf m) as HP. destruct HP as [HP _]. auto.
+ apply Maps.PTree.elements_keys_norepet. eassumption.
+ 2: { apply valueToPos_inj. apply unsigned_posToValue_le.
+ eapply pc_wf. intros. pose proof (HTL.mod_wf m) as HP. destruct HP as [HP _].
+ split. lia. apply HP. eassumption. eassumption.
+ apply unsigned_posToValue_le. eapply pc_wf. intros. pose proof (HTL.mod_wf m) as HP.
+ destruct HP as [HP _].
+ split. lia. apply HP. eassumption. eassumption. trivial.
+ }
+ apply Maps.PTree.elements_correct. eassumption. eassumption.
+
+ econstructor. econstructor.
+
+ eapply transl_list_correct.
+ intros. split. lia. pose proof (HTL.mod_wf m) as HP. destruct HP as [_ HP]. auto.
+ apply Maps.PTree.elements_keys_norepet. eassumption.
+ 2: { apply valueToPos_inj. apply unsigned_posToValue_le.
+ eapply pc_wf. intros. pose proof (HTL.mod_wf m) as HP. destruct HP as [HP _].
+ split. lia. apply HP. eassumption. eassumption.
+ apply unsigned_posToValue_le. eapply pc_wf. intros. pose proof (HTL.mod_wf m) as HP.
+ destruct HP as [HP _].
+ split. lia. apply HP. eassumption. eassumption. trivial.
+ }
+ apply Maps.PTree.elements_correct. eassumption. eassumption.
+
+ apply mis_stepp_decl. trivial. trivial. simpl. eassumption. auto.
+ rewrite valueToPos_posToValue. constructor; assumption. lia.
+
+ - econstructor; split. apply Smallstep.plus_one. apply step_finish. assumption. eassumption.
+ constructor; assumption.
+ - econstructor; split. apply Smallstep.plus_one. constructor.
+
+ constructor. constructor.
+ - inv H3. econstructor; split. apply Smallstep.plus_one. constructor. trivial.
+
+ apply match_state. assumption.
+ Qed.
+ Hint Resolve transl_step_correct : verilogproof.
+
+ Lemma transl_initial_states :
+ forall s1 : Smallstep.state (HTL.semantics prog),
+ Smallstep.initial_state (HTL.semantics prog) s1 ->
+ exists s2 : Smallstep.state (Verilog.semantics tprog),
+ Smallstep.initial_state (Verilog.semantics tprog) s2 /\ match_states s1 s2.
+ Proof.
+ induction 1.
+ econstructor; split. econstructor.
+ apply (Genv.init_mem_transf TRANSL); eauto.
+ rewrite symbols_preserved.
+ replace (AST.prog_main tprog) with (AST.prog_main prog); eauto.
+ symmetry; eapply Linking.match_program_main; eauto.
+ exploit function_ptr_translated; eauto. intros [tf [A B]].
+ inv B. eauto.
+ constructor.
+ Qed.
+ Hint Resolve transl_initial_states : verilogproof.
+
+ Lemma transl_final_states :
+ forall (s1 : Smallstep.state (HTL.semantics prog))
+ (s2 : Smallstep.state (Verilog.semantics tprog))
+ (r : Integers.Int.int),
+ match_states s1 s2 ->
+ Smallstep.final_state (HTL.semantics prog) s1 r ->
+ Smallstep.final_state (Verilog.semantics tprog) s2 r.
+ Proof.
+ intros. inv H0. inv H. inv H3. constructor. reflexivity.
+ Qed.
+ Hint Resolve transl_final_states : verilogproof.
Theorem transf_program_correct:
forward_simulation (HTL.semantics prog) (Verilog.semantics tprog).
- Admitted.
-
+ Proof.
+ eapply Smallstep.forward_simulation_plus; eauto with verilogproof.
+ apply senv_preserved.
+ Qed.
End CORRECTNESS.
-
diff --git a/src/translation/Veriloggenspec.v b/src/translation/Veriloggenspec.v
deleted file mode 100644
index 7dc632c..0000000
--- a/src/translation/Veriloggenspec.v
+++ /dev/null
@@ -1,17 +0,0 @@
-(*
- * 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/>.
- *)
diff --git a/src/verilog/Array.v b/src/verilog/Array.v
index f3e1cd7..be06541 100644
--- a/src/verilog/Array.v
+++ b/src/verilog/Array.v
@@ -29,7 +29,7 @@ Lemma list_set_spec1 {A : Type} :
forall l i (x : A),
i < length l -> nth_error (list_set i x l) i = Some x.
Proof.
- induction l; intros; destruct i; simplify; try lia; try reflexivity; firstorder.
+ induction l; intros; destruct i; crush; firstorder.
Qed.
Hint Resolve list_set_spec1 : array.
@@ -37,7 +37,7 @@ Lemma list_set_spec2 {A : Type} :
forall l i (x : A) d,
i < length l -> nth i (list_set i x l) d = x.
Proof.
- induction l; intros; destruct i; simplify; try lia; try reflexivity; firstorder.
+ induction l; intros; destruct i; crush; firstorder.
Qed.
Hint Resolve list_set_spec2 : array.
@@ -46,7 +46,7 @@ Lemma list_set_spec3 {A : Type} :
i1 <> i2 ->
nth_error (list_set i1 x l) i2 = nth_error l i2.
Proof.
- induction l; intros; destruct i1; destruct i2; simplify; try lia; try reflexivity; firstorder.
+ induction l; intros; destruct i1; destruct i2; crush; firstorder.
Qed.
Hint Resolve list_set_spec3 : array.
@@ -56,7 +56,7 @@ Lemma array_set_wf {A : Type} :
Proof.
induction l; intros; destruct i; auto.
- invert H; simplify; auto.
+ invert H; crush; auto.
Qed.
Definition array_set {A : Type} (i : nat) (x : A) (a : Array A) :=
@@ -72,7 +72,7 @@ Proof.
intros.
rewrite <- a.(arr_wf) in H.
- unfold array_set. simplify.
+ unfold array_set. crush.
eauto with array.
Qed.
Hint Resolve array_set_spec1 : array.
@@ -84,7 +84,7 @@ Proof.
intros.
rewrite <- a.(arr_wf) in H.
- unfold array_set. simplify.
+ unfold array_set. crush.
eauto with array.
Qed.
Hint Resolve array_set_spec2 : array.
@@ -93,7 +93,7 @@ Lemma array_set_len {A : Type} :
forall a i (x : A),
a.(arr_length) = (array_set i x a).(arr_length).
Proof.
- unfold array_set. simplify. reflexivity.
+ unfold array_set. crush.
Qed.
Definition array_get_error {A : Type} (i : nat) (a : Array A) : option A :=
@@ -104,7 +104,7 @@ Lemma array_get_error_equal {A : Type} :
a.(arr_contents) = b.(arr_contents) ->
array_get_error i a = array_get_error i b.
Proof.
- unfold array_get_error. congruence.
+ unfold array_get_error. crush.
Qed.
Lemma array_get_error_bound {A : Type} :
@@ -142,7 +142,7 @@ Proof.
unfold array_get_error.
unfold array_set.
- simplify.
+ crush.
eauto with array.
Qed.
@@ -180,18 +180,17 @@ Fixpoint list_repeat' {A : Type} (acc : list A) (a : A) (n : nat) : list A :=
Lemma list_repeat'_len {A : Type} : forall (a : A) n l,
length (list_repeat' l a n) = (n + Datatypes.length l)%nat.
Proof.
- induction n; intros; simplify; try reflexivity.
+ induction n; intros; crush; try reflexivity.
specialize (IHn (a :: l)).
rewrite IHn.
- simplify.
- lia.
+ crush.
Qed.
Lemma list_repeat'_app {A : Type} : forall (a : A) n l,
list_repeat' l a n = list_repeat' [] a n ++ l.
Proof.
- induction n; intros; simplify; try reflexivity.
+ induction n; intros; crush; try reflexivity.
pose proof IHn.
specialize (H (a :: l)).
@@ -207,7 +206,7 @@ Qed.
Lemma list_repeat'_head_tail {A : Type} : forall n (a : A),
a :: list_repeat' [] a n = list_repeat' [] a n ++ [a].
Proof.
- induction n; intros; simplify; try reflexivity.
+ induction n; intros; crush; try reflexivity.
rewrite list_repeat'_app.
replace (a :: list_repeat' [] a n ++ [a]) with (list_repeat' [] a n ++ [a] ++ [a]).
@@ -232,17 +231,17 @@ Proof.
intros.
unfold list_repeat.
rewrite list_repeat'_len.
- simplify. lia.
+ crush.
Qed.
Lemma dec_list_repeat_spec {A : Type} : forall n (a : A) a',
(forall x x' : A, {x' = x} + {~ x' = x}) ->
In a' (list_repeat a n) -> a' = a.
Proof.
- induction n; intros; simplify; try contradiction.
+ induction n; intros; crush.
unfold list_repeat in *.
- simplify.
+ crush.
rewrite list_repeat'_app in H.
pose proof (X a a').
@@ -278,14 +277,19 @@ Lemma list_repeat_lookup {A : Type} :
Proof.
induction n; intros.
- destruct i; simplify; lia.
+ destruct i; crush.
rewrite list_repeat_cons.
- destruct i; simplify; firstorder.
+ destruct i; crush; firstorder.
Qed.
Definition arr_repeat {A : Type} (a : A) (n : nat) : Array A := make_array (list_repeat a n).
+Lemma arr_repeat_length {A : Type} : forall n (a : A), arr_length (arr_repeat a n) = n.
+Proof.
+ unfold list_repeat. crush. apply list_repeat_len.
+Qed.
+
Fixpoint list_combine {A B C : Type} (f : A -> B -> C) (x : list A) (y : list B) : list C :=
match x, y with
| a :: t, b :: t' => f a b :: list_combine f t t'
@@ -295,9 +299,9 @@ Fixpoint list_combine {A B C : Type} (f : A -> B -> C) (x : list A) (y : list B)
Lemma list_combine_length {A B C : Type} (f : A -> B -> C) : forall (x : list A) (y : list B),
length (list_combine f x y) = min (length x) (length y).
Proof.
- induction x; intros; simplify; try reflexivity.
+ induction x; intros; crush.
- destruct y; simplify; auto.
+ destruct y; crush; auto.
Qed.
Definition combine {A B C : Type} (f : A -> B -> C) (x : Array A) (y : Array B) : Array C :=
@@ -310,15 +314,24 @@ Proof.
unfold combine.
unfold make_array.
- simplify.
+ crush.
rewrite <- (arr_wf x) in *.
rewrite <- (arr_wf y) in *.
- destruct (arr_contents x); simplify.
- - reflexivity.
- - destruct (arr_contents y); simplify.
- f_equal.
- rewrite list_combine_length.
- destruct (Min.min_dec (length l) (length l0)); congruence.
+ destruct (arr_contents x); destruct (arr_contents y); crush.
+ rewrite list_combine_length.
+ destruct (Min.min_dec (length l) (length l0)); congruence.
Qed.
+
+Ltac array :=
+ try match goal with
+ | [ |- context[arr_length (combine _ _ _)] ] =>
+ rewrite combine_length
+ | [ |- context[length (list_repeat _ _)] ] =>
+ rewrite list_repeat_len
+ | |- context[array_get_error _ (arr_repeat ?x _) = Some ?x] =>
+ unfold array_get_error, arr_repeat
+ | |- context[nth_error (list_repeat ?x _) _ = Some ?x] =>
+ apply list_repeat_lookup
+ end.
diff --git a/src/verilog/AssocMap.v b/src/verilog/AssocMap.v
index 5d531d5..c5cfa3f 100644
--- a/src/verilog/AssocMap.v
+++ b/src/verilog/AssocMap.v
@@ -16,7 +16,7 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-From coqup Require Import Coquplib Value.
+From coqup Require Import Coquplib ValueInt.
From compcert Require Import Maps.
Definition reg := positive.
@@ -192,7 +192,7 @@ Import AssocMapExt.
Definition assocmap := AssocMap.t value.
Definition find_assocmap (n : nat) : reg -> assocmap -> value :=
- get_default value (ZToValue n 0).
+ get_default value (ZToValue 0).
Definition empty_assocmap : assocmap := AssocMap.empty value.
diff --git a/src/verilog/HTL.v b/src/verilog/HTL.v
index a3623f0..b3d1442 100644
--- a/src/verilog/HTL.v
+++ b/src/verilog/HTL.v
@@ -17,13 +17,11 @@
*)
From Coq Require Import FSets.FMapPositive.
-From coqup Require Import Coquplib Value AssocMap Array.
+From coqup Require Import Coquplib ValueInt AssocMap Array.
From coqup Require Verilog.
From compcert Require Events Globalenvs Smallstep Integers Values.
From compcert Require Import Maps.
-Import HexNotationValue.
-
(** 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
@@ -38,6 +36,11 @@ Definition node := positive.
Definition datapath := PTree.t Verilog.stmnt.
Definition controllogic := PTree.t Verilog.stmnt.
+Definition map_well_formed {A : Type} (m : PTree.t A) : Prop :=
+ forall p0 : positive,
+ In p0 (map fst (Maps.PTree.elements m)) ->
+ Z.pos p0 <= Integers.Int.max_unsigned.
+
Record module: Type :=
mkmodule {
mod_params : list reg;
@@ -54,6 +57,7 @@ Record module: Type :=
mod_clk : reg;
mod_scldecls : AssocMap.t (option Verilog.io * Verilog.scl_decl);
mod_arrdecls : AssocMap.t (option Verilog.io * Verilog.arr_decl);
+ mod_wf : (map_well_formed mod_controllogic /\ map_well_formed mod_datapath);
}.
Definition fundef := AST.fundef module.
@@ -99,14 +103,15 @@ Inductive state : Type :=
Inductive step : genv -> state -> Events.trace -> state -> Prop :=
| step_module :
- forall g m st sf ctrl data ist
+ forall g m st sf ctrl data
asr asa
basr1 basa1 nasr1 nasa1
basr2 basa2 nasr2 nasa2
asr' asa'
- f stval pstval,
- asr!(m.(mod_st)) = Some ist ->
- valueToPos ist = st ->
+ f pstval,
+ asr!(mod_reset m) = Some (ZToValue 0) ->
+ asr!(mod_finish m) = Some (ZToValue 0) ->
+ asr!(m.(mod_st)) = Some (posToValue st) ->
m.(mod_controllogic)!st = Some ctrl ->
m.(mod_datapath)!st = Some data ->
Verilog.stmnt_runp f
@@ -115,6 +120,7 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop :=
ctrl
(Verilog.mkassociations basr1 nasr1)
(Verilog.mkassociations basa1 nasa1) ->
+ basr1!(m.(mod_st)) = Some (posToValue st) ->
Verilog.stmnt_runp f
(Verilog.mkassociations basr1 nasr1)
(Verilog.mkassociations basa1 nasa1)
@@ -123,26 +129,28 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop :=
(Verilog.mkassociations basa2 nasa2) ->
asr' = Verilog.merge_regs nasr2 basr2 ->
asa' = Verilog.merge_arrs nasa2 basa2 ->
- asr'!(m.(mod_st)) = Some stval ->
- valueToPos stval = pstval ->
+ asr'!(m.(mod_st)) = Some (posToValue pstval) ->
+ Z.pos pstval <= Integers.Int.max_unsigned ->
step g (State sf m st asr asa) Events.E0 (State sf m pstval asr' asa')
| step_finish :
forall g m st asr asa retval sf,
- asr!(m.(mod_finish)) = Some (1'h"1") ->
+ asr!(m.(mod_finish)) = Some (ZToValue 1) ->
asr!(m.(mod_return)) = Some retval ->
step g (State sf m st asr asa) Events.E0 (Returnstate sf retval)
| step_call :
forall g m args res,
step g (Callstate res m args) Events.E0
(State res m m.(mod_entrypoint)
- (AssocMap.set (mod_st m) (posToValue 32 m.(mod_entrypoint))
- (init_regs args m.(mod_params)))
+ (AssocMap.set (mod_reset m) (ZToValue 0)
+ (AssocMap.set (mod_finish m) (ZToValue 0)
+ (AssocMap.set (mod_st m) (posToValue m.(mod_entrypoint))
+ (init_regs args m.(mod_params)))))
(empty_stack m))
| step_return :
forall g m asr asa i r sf pc mst,
mst = mod_st m ->
step g (Returnstate (Stackframe r m pc asr asa :: sf) i) Events.E0
- (State sf m pc ((asr # mst <- (posToValue 32 pc)) # r <- i) asa).
+ (State sf m pc ((asr # mst <- (posToValue pc)) # r <- i) asa).
Hint Constructors step : htl.
Inductive initial_state (p: program): state -> Prop :=
diff --git a/src/verilog/PrintHTL.ml b/src/verilog/PrintHTL.ml
new file mode 100644
index 0000000..36fdd3c
--- /dev/null
+++ b/src/verilog/PrintHTL.ml
@@ -0,0 +1,81 @@
+(* -*- mode: tuareg -*-
+ * CoqUp: Verified high-level synthesis.
+ * Copyright (C) 2019-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/>.
+ *)
+
+open ValueInt
+open Datatypes
+open Camlcoq
+open AST
+open Clflags
+open Printf
+open Maps
+open AST
+open HTL
+open PrintAST
+open PrintVerilog
+open CoqupClflags
+
+let pstr pp = fprintf pp "%s"
+
+let reg pp r =
+ fprintf pp "x%d" (P.to_int r)
+
+let rec regs pp = function
+ | [] -> ()
+ | [r] -> reg pp r
+ | r1::rl -> fprintf pp "%a, %a" reg r1 regs rl
+
+let print_instruction pp (pc, i) =
+ fprintf pp "%5d:\t%s" pc (pprint_stmnt 0 i)
+
+let print_module pp id f =
+ fprintf pp "%s(%a) {\n" (extern_atom id) regs f.mod_params;
+ let datapath =
+ List.sort
+ (fun (pc1, _) (pc2, _) -> compare pc2 pc1)
+ (List.rev_map
+ (fun (pc, i) -> (P.to_int pc, i))
+ (PTree.elements f.mod_datapath)) in
+ let controllogic =
+ List.sort
+ (fun (pc1, _) (pc2, _) -> compare pc2 pc1)
+ (List.rev_map
+ (fun (pc, i) -> (P.to_int pc, i))
+ (PTree.elements f.mod_controllogic)) in
+ fprintf pp " datapath {\n";
+ List.iter (print_instruction pp) datapath;
+ fprintf pp " }\n\n controllogic {\n";
+ List.iter (print_instruction pp) controllogic;
+ fprintf pp " }\n}\n\n"
+
+let print_globdef pp (id, gd) =
+ match gd with
+ | Gfun(Internal f) -> print_module pp id f
+ | _ -> ()
+
+let print_program pp prog =
+ List.iter (print_globdef pp) prog.prog_defs
+
+let destination : string option ref = ref None
+
+let print_if prog =
+ match !destination with
+ | None -> ()
+ | Some f ->
+ let oc = open_out f in
+ print_program oc prog;
+ close_out oc
diff --git a/src/verilog/PrintVerilog.ml b/src/verilog/PrintVerilog.ml
index 6d10887..7f3eb29 100644
--- a/src/verilog/PrintVerilog.ml
+++ b/src/verilog/PrintVerilog.ml
@@ -17,7 +17,7 @@
*)
open Verilog
-open Value
+open ValueInt
open Datatypes
open Camlcoq
@@ -70,16 +70,30 @@ let unop = function
let register a = sprintf "reg_%d" (P.to_int a)
-let literal l = sprintf "%d'd%d" (Nat.to_int l.vsize) (Z.to_int (uvalueToZ l))
+let literal l = sprintf "32'd%d" (Z.to_int (uvalueToZ l))
-let rec pprint_expr = function
+let literal_int i = sprintf "32'd%d" i
+
+let byte n s = sprintf "reg_%d[%d:%d]" (P.to_int s) (7 + n * 8) (n * 8)
+
+
+let rec pprint_expr =
+ let array_byte r i = function
+ | 0 -> concat [register r; "["; pprint_expr i; "]"]
+ | n -> concat [register r; "["; pprint_expr i; " + "; literal_int n; "][7:0]"]
+ in function
| Vlit l -> literal l
| Vvar s -> register s
+ | Vvarb0 s -> byte 0 s
+ | Vvarb1 s -> byte 1 s
+ | Vvarb2 s -> byte 2 s
+ | Vvarb3 s -> byte 3 s
| Vvari (s, i) -> concat [register s; "["; pprint_expr i; "]"]
| Vinputvar s -> register s
| Vunop (u, e) -> concat ["("; unop u; pprint_expr e; ")"]
- | Vbinop (op, a, b) -> concat ["("; pprint_binop (pprint_expr a) (pprint_expr b) op; ")"]
+ | Vbinop (op, a, b) -> concat [pprint_binop (pprint_expr a) (pprint_expr b) op]
| Vternary (c, t, f) -> concat ["("; pprint_expr c; " ? "; pprint_expr t; " : "; pprint_expr f; ")"]
+ | Vload (s, i) -> concat ["{"; array_byte s i 3; ", "; array_byte s i 2; ", "; array_byte s i 1; ", "; array_byte s i 0; "}"]
let rec pprint_stmnt i =
let pprint_case (e, s) = concat [ indent (i + 1); pprint_expr e; ": begin\n"; pprint_stmnt (i + 2) s;
diff --git a/src/verilog/PrintVerilog.mli b/src/verilog/PrintVerilog.mli
index 6544e52..5fd8fe9 100644
--- a/src/verilog/PrintVerilog.mli
+++ b/src/verilog/PrintVerilog.mli
@@ -16,8 +16,10 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-val print_value : out_channel -> Value.value -> unit
+val pprint_stmnt : int -> Verilog.stmnt -> string
+
+val print_value : out_channel -> ValueInt.value -> unit
val print_program : bool -> out_channel -> Verilog.program -> unit
-val print_result : out_channel -> (BinNums.positive * Value.value) list -> unit
+val print_result : out_channel -> (BinNums.positive * ValueInt.value) list -> unit
diff --git a/src/verilog/Value.v b/src/verilog/Value.v
index 8ba5138..af2b822 100644
--- a/src/verilog/Value.v
+++ b/src/verilog/Value.v
@@ -21,6 +21,7 @@ From bbv Require Import Word.
From bbv Require HexNotation WordScope.
From Coq Require Import ZArith.ZArith FSets.FMapPositive Lia.
From compcert Require Import lib.Integers common.Values.
+From coqup Require Import Coquplib.
(* end hide *)
(** * Value
@@ -39,6 +40,8 @@ Record value : Type :=
vword: word vsize
}.
+Search N.of_nat.
+
(** ** Value conversions
Various conversions to different number types such as [N], [Z], [positive] and
@@ -85,9 +88,18 @@ Definition intToValue (i : Integers.int) : value :=
Definition valueToInt (i : value) : Integers.int :=
Int.repr (uvalueToZ i).
+Definition ptrToValue (i : Integers.ptrofs) : value :=
+ ZToValue Ptrofs.wordsize (Ptrofs.unsigned i).
+
+Definition valueToPtr (i : value) : Integers.ptrofs :=
+ Ptrofs.repr (uvalueToZ i).
+
Definition valToValue (v : Values.val) : option value :=
match v with
| Values.Vint i => Some (intToValue i)
+ | Values.Vptr b off => if Z.eqb (Z.modulo (uvalueToZ (ptrToValue off)) 4) 0%Z
+ then Some (ptrToValue off)
+ else None
| Values.Vundef => Some (ZToValue 32 0%Z)
| _ => None
end.
@@ -301,8 +313,8 @@ Inductive val_value_lessdef: val -> value -> Prop :=
val_value_lessdef (Vint i) v'
| val_value_lessdef_ptr:
forall b off v',
- off = Ptrofs.repr (valueToZ v') ->
- (Z.modulo (valueToZ v') 4) = 0%Z ->
+ off = valueToPtr v' ->
+ (Z.modulo (uvalueToZ v') 4) = 0%Z ->
val_value_lessdef (Vptr b off) v'
| lessdef_undef: forall v, val_value_lessdef Vundef v.
@@ -379,6 +391,41 @@ Proof.
apply Z.lt_le_pred in H. apply H.
Qed.
+Lemma valueToPtr_ptrToValue :
+ forall v,
+ valueToPtr (ptrToValue v) = v.
+Proof.
+ intros.
+ unfold valueToPtr, ptrToValue. rewrite uvalueToZ_ZToValue. auto using Ptrofs.repr_unsigned.
+ split. apply Ptrofs.unsigned_range_2.
+ assert ((Ptrofs.unsigned v <= Ptrofs.max_unsigned)%Z) by apply Ptrofs.unsigned_range_2.
+ apply Z.lt_le_pred in H. apply H.
+Qed.
+
+Lemma intToValue_valueToInt :
+ forall v,
+ vsize v = 32%nat ->
+ intToValue (valueToInt v) = v.
+Proof.
+ intros. unfold valueToInt, intToValue. rewrite Int.unsigned_repr_eq.
+ unfold ZToValue, uvalueToZ. unfold Int.modulus. unfold Int.wordsize. unfold Wordsize_32.wordsize.
+ pose proof (uwordToZ_bound (vword v)).
+ rewrite Z.mod_small. rewrite <- H. rewrite ZToWord_uwordToZ. destruct v; auto.
+ rewrite <- H. rewrite two_power_nat_equiv. apply H0.
+Qed.
+
+Lemma ptrToValue_valueToPtr :
+ forall v,
+ vsize v = 32%nat ->
+ ptrToValue (valueToPtr v) = v.
+Proof.
+ intros. unfold valueToPtr, ptrToValue. rewrite Ptrofs.unsigned_repr_eq.
+ unfold ZToValue, uvalueToZ. unfold Ptrofs.modulus. unfold Ptrofs.wordsize. unfold Wordsize_Ptrofs.wordsize.
+ pose proof (uwordToZ_bound (vword v)).
+ rewrite Z.mod_small. rewrite <- H. rewrite ZToWord_uwordToZ. destruct v; auto.
+ rewrite <- H. rewrite two_power_nat_equiv. apply H0.
+Qed.
+
Lemma valToValue_lessdef :
forall v v',
valToValue v = Some v' ->
@@ -388,6 +435,10 @@ Proof.
destruct v; try discriminate; constructor.
unfold valToValue in H. inversion H.
symmetry. apply valueToInt_intToValue.
+ inv H. destruct (uvalueToZ (ptrToValue i) mod 4 =? 0); try discriminate.
+ inv H1. symmetry. apply valueToPtr_ptrToValue.
+ inv H. destruct (uvalueToZ (ptrToValue i) mod 4 =? 0) eqn:?; try discriminate.
+ inv H1. apply Z.eqb_eq. apply Heqb0.
Qed.
Lemma boolToValue_ValueToBool :
@@ -415,6 +466,10 @@ Proof.
rewrite ZToWord_plus; auto.
Qed.
+Lemma ZToValue_eq :
+ forall w1,
+ (mkvalue 32 w1) = (ZToValue 32 (wordToZ w1)). Abort.
+
Lemma wordsize_32 :
Int.wordsize = 32%nat.
Proof. auto. Qed.
@@ -428,6 +483,28 @@ Proof.
rewrite Int.repr_unsigned. auto. rewrite wordsize_32. omega.
Qed.
+(*Lemma intadd_vplus2 :
+ forall v1 v2 EQ,
+ vsize v1 = 32%nat ->
+ Int.add (valueToInt v1) (valueToInt v2) = valueToInt (vplus v1 v2 EQ).
+Proof.
+ intros. unfold Int.add, valueToInt, intToValue. repeat (rewrite Int.unsigned_repr).
+ rewrite (@vadd_vplus v1 v2 EQ). trivial.
+ unfold uvalueToZ. Search word "bound". pose proof (@uwordToZ_bound (vsize v2) (vword v2)).
+ rewrite H in EQ. rewrite <- EQ in H0 at 3.*)
+ (*rewrite zadd_vplus3. trivia*)
+
+Lemma valadd_vplus :
+ forall v1 v2 v1' v2' v v' EQ,
+ val_value_lessdef v1 v1' ->
+ val_value_lessdef v2 v2' ->
+ Val.add v1 v2 = v ->
+ vplus v1' v2' EQ = v' ->
+ val_value_lessdef v v'.
+Proof.
+ intros. inv H; inv H0; constructor; simplify.
+ Abort.
+
Lemma zsub_vminus :
forall sz z1 z2,
(sz > 0)%nat ->
@@ -452,17 +529,25 @@ Proof.
unfold wordBin. repeat (rewrite wordToN_NToWord_2); auto.
Qed.
-(*Lemma ZToValue_valueToNat :
+Lemma ZToValue_valueToNat :
forall x sz,
- sz > 0 ->
- (x < 2^(Z.of_nat sz))%Z ->
+ (sz > 0)%nat ->
+ (0 <= x < 2^(Z.of_nat sz))%Z ->
valueToNat (ZToValue sz x) = Z.to_nat x.
Proof.
- destruct x; intros; unfold ZToValue, valueToNat; simpl.
+ destruct x; intros; unfold ZToValue, valueToNat; crush.
- rewrite wzero'_def. apply wordToNat_wzero.
- rewrite posToWord_nat. rewrite wordToNat_natToWord_2. trivial.
- unfold Z.of_nat in *. destruct sz eqn:?. omega. simpl in H0.
- rewrite <- Pos2Z.inj_pow_pos in H0. Search (Z.pos _ < Z.pos _)%Z.
- Search Pos.to_nat (_ < _). (* Pos2Nat.inj_lt *)
- Search "inj" positive nat.
-*)
+ clear H1.
+ lazymatch goal with
+ | [ H : context[(_ < ?x)%Z] |- _ ] => replace x with (Z.of_nat (Z.to_nat x)) in H
+ end.
+ 2: { apply Z2Nat.id; apply Z.pow_nonneg; lia. }
+
+ rewrite Z2Nat.inj_pow in H2; crush.
+ replace (Pos.to_nat 2) with 2%nat in H2 by reflexivity.
+ rewrite Nat2Z.id in H2.
+ rewrite <- positive_nat_Z in H2.
+ apply Nat2Z.inj_lt in H2.
+ assumption.
+Qed.
diff --git a/src/verilog/ValueInt.v b/src/verilog/ValueInt.v
new file mode 100644
index 0000000..151feef
--- /dev/null
+++ b/src/verilog/ValueInt.v
@@ -0,0 +1,162 @@
+(*
+ * 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/>.
+ *)
+
+(* begin hide *)
+From bbv Require Import Word.
+From bbv Require HexNotation WordScope.
+From Coq Require Import ZArith.ZArith FSets.FMapPositive Lia.
+From compcert Require Import lib.Integers common.Values.
+From coqup Require Import Coquplib.
+(* end hide *)
+
+(** * Value
+
+A [value] is a bitvector with a specific size. We are using the implementation
+of the bitvector by mit-plv/bbv, because it has many theorems that we can reuse.
+However, we need to wrap it with an [Inductive] so that we can specify and match
+on the size of the [value]. This is necessary so that we can easily store
+[value]s of different sizes in a list or in a map.
+
+Using the default [word], this would not be possible, as the size is part of the type. *)
+
+Definition value : Type := int.
+
+(** ** Value conversions
+
+Various conversions to different number types such as [N], [Z], [positive] and
+[int], where the last one is a theory of integers of powers of 2 in CompCert. *)
+
+Definition valueToNat (v : value) : nat :=
+ Z.to_nat (Int.unsigned v).
+
+Definition natToValue (n : nat) : value :=
+ Int.repr (Z.of_nat n).
+
+Definition valueToN (v : value) : N :=
+ Z.to_N (Int.unsigned v).
+
+Definition NToValue (n : N) : value :=
+ Int.repr (Z.of_N n).
+
+Definition ZToValue (z : Z) : value :=
+ Int.repr z.
+
+Definition valueToZ (v : value) : Z :=
+ Int.signed v.
+
+Definition uvalueToZ (v : value) : Z :=
+ Int.unsigned v.
+
+Definition posToValue (p : positive) : value :=
+ Int.repr (Z.pos p).
+
+Definition valueToPos (v : value) : positive :=
+ Z.to_pos (Int.unsigned v).
+
+Definition intToValue (i : Integers.int) : value := i.
+
+Definition valueToInt (i : value) : Integers.int := i.
+
+Definition ptrToValue (i : ptrofs) : value := Ptrofs.to_int i.
+
+Definition valueToPtr (i : value) : Integers.ptrofs :=
+ Ptrofs.of_int i.
+
+Definition valToValue (v : Values.val) : option value :=
+ match v with
+ | Values.Vint i => Some (intToValue i)
+ | Values.Vptr b off => Some (ptrToValue off)
+ | Values.Vundef => Some (ZToValue 0%Z)
+ | _ => None
+ end.
+
+(** Convert a [value] to a [bool], so that choices can be made based on the
+result. This is also because comparison operators will give back [value] instead
+of [bool], so if they are in a condition, they will have to be converted before
+they can be used. *)
+
+Definition valueToBool (v : value) : bool :=
+ if Z.eqb (uvalueToZ v) 0 then false else true.
+
+Definition boolToValue (b : bool) : value :=
+ natToValue (if b then 1 else 0).
+
+(** ** Arithmetic operations *)
+
+Definition unify_word (sz1 sz2 : nat) (w1 : word sz2): sz1 = sz2 -> word sz1.
+intros; subst; assumption. Defined.
+
+Lemma unify_word_unfold :
+ forall sz w,
+ unify_word sz sz w eq_refl = w.
+Proof. auto. Qed.
+
+Inductive val_value_lessdef: val -> value -> Prop :=
+| val_value_lessdef_int:
+ forall i v',
+ i = valueToInt v' ->
+ val_value_lessdef (Vint i) v'
+| val_value_lessdef_ptr:
+ forall b off v',
+ off = valueToPtr v' ->
+ val_value_lessdef (Vptr b off) v'
+| lessdef_undef: forall v, val_value_lessdef Vundef v.
+
+Inductive opt_val_value_lessdef: option val -> value -> Prop :=
+| opt_lessdef_some:
+ forall v v', val_value_lessdef v v' -> opt_val_value_lessdef (Some v) v'
+| opt_lessdef_none: forall v, opt_val_value_lessdef None v.
+
+Lemma valueToZ_ZToValue :
+ forall z,
+ (Int.min_signed <= z <= Int.max_signed)%Z ->
+ valueToZ (ZToValue z) = z.
+Proof. auto using Int.signed_repr. Qed.
+
+Lemma uvalueToZ_ZToValue :
+ forall z,
+ (0 <= z <= Int.max_unsigned)%Z ->
+ uvalueToZ (ZToValue z) = z.
+Proof. auto using Int.unsigned_repr. Qed.
+
+Lemma valueToPos_posToValue :
+ forall v,
+ 0 <= Z.pos v <= Int.max_unsigned ->
+ valueToPos (posToValue v) = v.
+Proof.
+ unfold valueToPos, posToValue.
+ intros. rewrite Int.unsigned_repr.
+ apply Pos2Z.id. assumption.
+Qed.
+
+Lemma valueToInt_intToValue :
+ forall v,
+ valueToInt (intToValue v) = v.
+Proof. auto. Qed.
+
+Lemma valToValue_lessdef :
+ forall v v',
+ valToValue v = Some v' ->
+ val_value_lessdef v v'.
+Proof.
+ intros.
+ destruct v; try discriminate; constructor.
+ unfold valToValue in H. inversion H.
+ unfold valueToInt. unfold intToValue in H1. auto.
+ inv H. symmetry. unfold valueToPtr, ptrToValue. apply Ptrofs.of_int_to_int. trivial.
+Qed.
diff --git a/src/verilog/Verilog.v b/src/verilog/Verilog.v
index 555ddbd..7ede80c 100644
--- a/src/verilog/Verilog.v
+++ b/src/verilog/Verilog.v
@@ -29,11 +29,9 @@ Require Import Lia.
Import ListNotations.
-From coqup Require Import common.Coquplib common.Show verilog.Value AssocMap Array.
-From compcert Require Integers Events.
-From compcert Require Import Errors Smallstep Globalenvs.
-
-Import HexNotationValue.
+From coqup Require Import common.Coquplib common.Show verilog.ValueInt IntegerExtra AssocMap Array.
+From compcert Require Events.
+From compcert Require Import Integers Errors Smallstep Globalenvs.
Local Open Scope assocmap.
@@ -80,7 +78,7 @@ Definition merge_arrs (new : assocmap_arr) (old : assocmap_arr) : assocmap_arr :
Definition arr_assocmap_lookup (a : assocmap_arr) (r : reg) (i : nat) : option value :=
match a ! r with
| None => None
- | Some arr => Some (Option.default (NToValue 32 0) (Option.join (array_get_error i arr)))
+ | Some arr => Some (Option.default (NToValue 0) (Option.join (array_get_error i arr)))
end.
Definition arr_assocmap_set (r : reg) (i : nat) (v : value) (a : assocmap_arr) : assocmap_arr :=
@@ -156,16 +154,21 @@ Inductive unop : Type :=
(** ** Expressions *)
Inductive expr : Type :=
-| Vlit : value -> expr
-| Vvar : reg -> expr
-| Vvari : reg -> expr -> expr
+| Vlit : value -> expr (** literal *)
+| Vvar : reg -> expr (** reg *)
+| Vvarb0 : reg -> expr (** 1st byte projection of reg *)
+| Vvarb1 : reg -> expr
+| Vvarb2 : reg -> expr
+| Vvarb3 : reg -> expr
+| Vvari : reg -> expr -> expr (** array *)
| Vinputvar : reg -> expr
| Vbinop : binop -> expr -> expr -> expr
| Vunop : unop -> expr -> expr
-| Vternary : expr -> expr -> expr -> expr.
+| Vternary : expr -> expr -> expr -> expr
+| Vload : reg -> expr -> expr. (** 4-byte concatenation load *)
-Definition posToExpr (sz : nat) (p : positive) : expr :=
- Vlit (posToValue sz p).
+Definition posToExpr (p : positive) : expr :=
+ Vlit (posToValue p).
(** ** Statements *)
@@ -245,12 +248,9 @@ Definition program := AST.program fundef unit.
(** Convert a [positive] to an expression directly, setting it to the right
size. *)
Definition posToLit (p : positive) : expr :=
- Vlit (posToValueAuto p).
-
-Coercion Vlit : value >-> expr.
-Coercion Vvar : reg >-> expr.
+ Vlit (posToValue p).
-Definition fext := assocmap.
+Definition fext := unit.
Definition fextclk := nat -> fext.
(** ** State
@@ -298,81 +298,98 @@ Inductive state : Type :=
(m : module)
(args : list value), state.
-Definition binop_run (op : binop) : forall v1 v2 : value, vsize v1 = vsize v2 -> value :=
+Definition binop_run (op : binop) (v1 v2 : value) : option value :=
match op with
- | Vadd => vplus
- | Vsub => vminus
- | Vmul => vmul
- | Vdiv => vdivs
- | Vdivu => vdiv
- | Vmod => vmods
- | Vmodu => vmod
- | Vlt => vlts
- | Vltu => vlt
- | Vgt => vgts
- | Vgtu => vgt
- | Vle => vles
- | Vleu => vle
- | Vge => vges
- | Vgeu => vge
- | Veq => veq
- | Vne => vne
- | Vand => vand
- | Vor => vor
- | Vxor => vxor
- | Vshl => vshl
- | Vshr => vshr
- | Vshru => vshr (* FIXME: should not be the same operation. *)
+ | Vadd => Some (Int.add v1 v2)
+ | Vsub => Some (Int.sub v1 v2)
+ | Vmul => Some (Int.mul v1 v2)
+ | Vdiv => if Int.eq v2 Int.zero
+ || Int.eq v1 (Int.repr Int.min_signed) && Int.eq v2 Int.mone
+ then None
+ else Some (Int.divs v1 v2)
+ | Vdivu => if Int.eq v2 Int.zero then None else Some (Int.divu v1 v2)
+ | Vmod => if Int.eq v2 Int.zero
+ || Int.eq v1 (Int.repr Int.min_signed) && Int.eq v2 Int.mone
+ then None
+ else Some (Int.mods v1 v2)
+ | Vmodu => if Int.eq v2 Int.zero then None else Some (Int.modu v1 v2)
+ | Vlt => Some (boolToValue (Int.lt v1 v2))
+ | Vltu => Some (boolToValue (Int.ltu v1 v2))
+ | Vgt => Some (boolToValue (Int.lt v2 v1))
+ | Vgtu => Some (boolToValue (Int.ltu v2 v1))
+ | Vle => Some (boolToValue (negb (Int.lt v2 v1)))
+ | Vleu => Some (boolToValue (negb (Int.ltu v2 v1)))
+ | Vge => Some (boolToValue (negb (Int.lt v1 v2)))
+ | Vgeu => Some (boolToValue (negb (Int.ltu v1 v2)))
+ | Veq => Some (boolToValue (Int.eq v1 v2))
+ | Vne => Some (boolToValue (negb (Int.eq v1 v2)))
+ | Vand => Some (Int.and v1 v2)
+ | Vor => Some (Int.or v1 v2)
+ | Vxor => Some (Int.xor v1 v2)
+ | Vshl => Some (Int.shl v1 v2)
+ | Vshr => Some (Int.shr v1 v2)
+ | Vshru => Some (Int.shru v1 v2)
end.
-Definition unop_run (op : unop) : value -> value :=
+Definition unop_run (op : unop) (v1 : value) : value :=
match op with
- | Vneg => vnot
- | Vnot => vbitneg
+ | Vneg => Int.neg v1
+ | Vnot => Int.not v1
end.
Inductive expr_runp : fext -> assocmap -> assocmap_arr -> expr -> value -> Prop :=
| erun_Vlit :
- forall fext reg stack v,
- expr_runp fext reg stack (Vlit v) v
+ forall fext asr asa v,
+ expr_runp fext asr asa (Vlit v) v
| erun_Vvar :
- forall fext reg stack v r,
- reg#r = v ->
- expr_runp fext reg stack (Vvar r) v
+ forall fext asr asa v r,
+ asr#r = v ->
+ expr_runp fext asr asa (Vvar r) v
+ | erun_Vvarb0 :
+ forall fext asr asa v r,
+ asr#r = v ->
+ expr_runp fext asr asa (Vvarb0 r) (IntExtra.ibyte0 v)
+ | erun_Vvarb1 :
+ forall fext asr asa v r,
+ asr#r = v ->
+ expr_runp fext asr asa (Vvarb1 r) (IntExtra.ibyte1 v)
+ | erun_Vvarb2 :
+ forall fext asr asa v r,
+ asr#r = v ->
+ expr_runp fext asr asa (Vvarb2 r) (IntExtra.ibyte2 v)
+ | erun_Vvarb3 :
+ forall fext asr asa v r,
+ asr#r = v ->
+ expr_runp fext asr asa (Vvarb3 r) (IntExtra.ibyte3 v)
| erun_Vvari :
- forall fext reg stack v iexp i r,
- expr_runp fext reg stack iexp i ->
- arr_assocmap_lookup stack r (valueToNat i) = Some v ->
- expr_runp fext reg stack (Vvari r iexp) v
- | erun_Vinputvar :
- forall fext reg stack r v,
- fext!r = Some v ->
- expr_runp fext reg stack (Vinputvar r) v
+ forall fext asr asa v iexp i r,
+ expr_runp fext asr asa iexp i ->
+ arr_assocmap_lookup asa r (valueToNat i) = Some v ->
+ expr_runp fext asr asa (Vvari r iexp) v
| erun_Vbinop :
- forall fext reg stack op l r lv rv oper EQ resv,
- expr_runp fext reg stack l lv ->
- expr_runp fext reg stack r rv ->
- oper = binop_run op ->
- resv = oper lv rv EQ ->
- expr_runp fext reg stack (Vbinop op l r) resv
+ forall fext asr asa op l r lv rv resv,
+ expr_runp fext asr asa l lv ->
+ expr_runp fext asr asa r rv ->
+ Some resv = binop_run op lv rv ->
+ expr_runp fext asr asa (Vbinop op l r) resv
| erun_Vunop :
- forall fext reg stack u vu op oper resv,
- expr_runp fext reg stack u vu ->
+ forall fext asr asa u vu op oper resv,
+ expr_runp fext asr asa u vu ->
oper = unop_run op ->
resv = oper vu ->
- expr_runp fext reg stack (Vunop op u) resv
+ expr_runp fext asr asa (Vunop op u) resv
| erun_Vternary_true :
- forall fext reg stack c ts fs vc vt,
- expr_runp fext reg stack c vc ->
- expr_runp fext reg stack ts vt ->
+ forall fext asr asa c ts fs vc vt,
+ expr_runp fext asr asa c vc ->
+ expr_runp fext asr asa ts vt ->
valueToBool vc = true ->
- expr_runp fext reg stack (Vternary c ts fs) vt
+ expr_runp fext asr asa (Vternary c ts fs) vt
| erun_Vternary_false :
- forall fext reg stack c ts fs vc vf,
- expr_runp fext reg stack c vc ->
- expr_runp fext reg stack fs vf ->
+ forall fext asr asa c ts fs vc vf,
+ expr_runp fext asr asa c vc ->
+ expr_runp fext asr asa fs vf ->
valueToBool vc = false ->
- expr_runp fext reg stack (Vternary c ts fs) vf.
+ expr_runp fext asr asa (Vternary c ts fs) vf.
Hint Constructors expr_runp : verilog.
Definition handle_opt {A : Type} (err : errmsg) (val : option A)
@@ -391,11 +408,11 @@ 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 :=
+(*Definition access_fext (f : fext) (r : reg) : res value :=
match AssocMap.get r f with
| Some v => OK v
- | _ => OK (ZToValue 1 0)
- end.
+ | _ => OK (ZToValue 0)
+ end.*)
(* TODO FIX Vvar case without default *)
(*Fixpoint expr_run (assoc : assocmap) (e : expr)
@@ -432,8 +449,8 @@ Inductive location : Type :=
| LocArray (_ : reg) (_ : nat).
Inductive location_is : fext -> assocmap -> assocmap_arr -> expr -> location -> Prop :=
-| Base : forall f asr asa r, location_is f asr asa (Vvar r) (LocReg r)
-| Indexed : forall f asr asa r iexp iv,
+| Reg : forall f asr asa r, location_is f asr asa (Vvar r) (LocReg r)
+| RegIndexed : forall f asr asa r iexp iv,
expr_runp f asr asa iexp iv ->
location_is f asr asa (Vvari r iexp) (LocArray r (valueToNat iv)).
@@ -519,7 +536,7 @@ Inductive stmnt_runp: fext -> reg_associations -> arr_associations ->
forall f st1 st2 asr0 asa0 asr1 asa1 asr2 asa2,
stmnt_runp f asr0 asa0 st1 asr1 asa1 ->
stmnt_runp f asr1 asa1 st2 asr2 asa2 ->
- stmnt_runp f asr0 asa1 (Vseq st1 st2) asr2 asa2
+ stmnt_runp f asr0 asa0 (Vseq st1 st2) asr2 asa2
| stmnt_runp_Vcond_true:
forall asr0 asa0 asr1 asa1 f c vc stt stf,
expr_runp f asr0.(assoc_blocking) asa0.(assoc_blocking) c vc ->
@@ -530,7 +547,7 @@ Inductive stmnt_runp: fext -> reg_associations -> arr_associations ->
forall asr0 asa0 asr1 asa1 f c vc stt stf,
expr_runp f asr0.(assoc_blocking) asa0.(assoc_blocking) c vc ->
valueToBool vc = false ->
- stmnt_runp f asr0 asa0 stt asr1 asa1 ->
+ stmnt_runp f asr0 asa0 stf asr1 asa1 ->
stmnt_runp f asr0 asa0 (Vcond c stt stf) asr1 asa1
| stmnt_runp_Vcase_nomatch:
forall e ve asr0 asa0 f asr1 asa1 me mve sc cs def,
@@ -599,7 +616,7 @@ Inductive mi_stepp : fext -> reg_associations -> arr_associations ->
module_item -> reg_associations -> arr_associations -> Prop :=
| mi_stepp_Valways :
forall f sr0 sa0 st sr1 sa1 c,
- stmnt_runp f sr0 sa0 st sr1 sa0 ->
+ stmnt_runp f sr0 sa0 st sr1 sa1 ->
mi_stepp f sr0 sa0 (Valways c st) sr1 sa1
| mi_stepp_Valways_ff :
forall f sr0 sa0 st sr1 sa1 c,
@@ -610,8 +627,8 @@ Inductive mi_stepp : fext -> reg_associations -> arr_associations ->
stmnt_runp f sr0 sa0 st sr1 sa1 ->
mi_stepp f sr0 sa0 (Valways_comb c st) sr1 sa1
| mi_stepp_Vdecl :
- forall f sr sa lhs rhs io,
- mi_stepp f sr sa (Vdeclaration (Vdecl io lhs rhs)) sr sa.
+ forall f sr sa d,
+ mi_stepp f sr sa (Vdeclaration d) sr sa.
Hint Constructors mi_stepp : verilog.
Inductive mis_stepp : fext -> reg_associations -> arr_associations ->
@@ -650,11 +667,11 @@ Fixpoint mi_run (f : fextclk) (assoc : assocmap) (m : list module_item) (n : nat
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 :=
+(*Definition initial_fextclk (m : module) : fextclk :=
fun x => match x with
- | S O => (AssocMap.set (mod_reset m) (ZToValue 1 1) empty_assocmap)
- | _ => (AssocMap.set (mod_reset m) (ZToValue 1 0) empty_assocmap)
- end.
+ | S O => (AssocMap.set (mod_reset m) (ZToValue 1) empty_assocmap)
+ | _ => (AssocMap.set (mod_reset m) (ZToValue 0) empty_assocmap)
+ end.*)
(*Definition module_run (n : nat) (m : module) : res assocmap :=
mi_run (initial_fextclk m) empty_assocmap (mod_body m) n.*)
@@ -713,6 +730,8 @@ Definition empty_stack (m : module) : assocmap_arr :=
Inductive step : genv -> state -> Events.trace -> state -> Prop :=
| step_module :
forall asr asa asr' asa' basr1 nasr1 basa1 nasa1 f stval pstval m sf st g ist,
+ asr!(m.(mod_reset)) = Some (ZToValue 0) ->
+ asr!(m.(mod_finish)) = Some (ZToValue 0) ->
asr!(m.(mod_st)) = Some ist ->
valueToPos ist = st ->
mis_stepp f (mkassociations asr empty_assocmap)
@@ -727,22 +746,23 @@ Inductive step : genv -> state -> Events.trace -> state -> Prop :=
step g (State sf m st asr asa) Events.E0 (State sf m pstval asr' asa')
| step_finish :
forall asr asa sf retval m st g,
- asr!(m.(mod_finish)) = Some (1'h"1") ->
+ asr!(m.(mod_finish)) = Some (ZToValue 1) ->
asr!(m.(mod_return)) = Some retval ->
step g (State sf m st asr asa) Events.E0 (Returnstate sf retval)
| step_call :
forall g res m args,
step g (Callstate res m args) Events.E0
(State res m m.(mod_entrypoint)
- (AssocMap.set m.(mod_st) (posToValue 32 m.(mod_entrypoint))
- (init_params args m.(mod_args)))
+ (AssocMap.set (mod_reset m) (ZToValue 0)
+ (AssocMap.set (mod_finish m) (ZToValue 0)
+ (AssocMap.set m.(mod_st) (posToValue m.(mod_entrypoint))
+ (init_params args m.(mod_args)))))
(empty_stack m))
| step_return :
forall g m asr i r sf pc mst asa,
mst = mod_st m ->
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)).
+ (State sf m pc ((asr # mst <- (posToValue pc)) # r <- i) asa).
Hint Constructors step : verilog.
Inductive initial_state (p: program): state -> Prop :=
@@ -761,3 +781,138 @@ Inductive final_state : state -> Integers.int -> Prop :=
Definition semantics (m : program) :=
Smallstep.Semantics step (initial_state m) final_state
(Globalenvs.Genv.globalenv m).
+
+Lemma expr_runp_determinate :
+ forall f e asr asa v,
+ expr_runp f asr asa e v ->
+ forall v',
+ expr_runp f asr asa e v' ->
+ v' = v.
+Proof.
+ induction e; intros;
+
+ repeat (try match goal with
+ | [ H : expr_runp _ _ _ (Vlit _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vvar _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vvarb0 _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vvarb1 _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vvarb2 _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vvarb3 _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vvari _ _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vinputvar _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vbinop _ _ _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vunop _ _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vternary _ _ _) _ |- _ ] => invert H
+ | [ H : expr_runp _ _ _ (Vload _ _) _ |- _ ] => invert H
+
+ | [ H1 : forall asr asa v, expr_runp _ asr asa ?e v -> _,
+ H2 : expr_runp _ _ _ ?e _ |- _ ] =>
+ learn (H1 _ _ _ H2)
+ | [ H1 : forall v, expr_runp _ ?asr ?asa ?e v -> _,
+ H2 : expr_runp _ ?asr ?asa ?e _ |- _ ] =>
+ learn (H1 _ H2)
+ end; crush).
+Qed.
+Hint Resolve expr_runp_determinate : verilog.
+
+Lemma location_is_determinate :
+ forall f asr asa e l,
+ location_is f asr asa e l ->
+ forall l',
+ location_is f asr asa e l' ->
+ l' = l.
+Proof.
+ induction 1; intros;
+
+ repeat (try match goal with
+ | [ H : location_is _ _ _ _ _ |- _ ] => invert H
+ | [ H1 : expr_runp _ ?asr ?asa ?e _,
+ H2 : expr_runp _ ?asr ?asa ?e _ |- _ ] =>
+ learn (expr_runp_determinate H1 H2)
+ end; crush).
+Qed.
+
+Lemma stmnt_runp_determinate :
+ forall f s asr0 asa0 asr1 asa1,
+ stmnt_runp f asr0 asa0 s asr1 asa1 ->
+ forall asr1' asa1',
+ stmnt_runp f asr0 asa0 s asr1' asa1' ->
+ asr1' = asr1 /\ asa1' = asa1.
+ induction 1; intros;
+
+ repeat (try match goal with
+ | [ H : stmnt_runp _ _ _ (Vseq _ _) _ _ |- _ ] => invert H
+ | [ H : stmnt_runp _ _ _ (Vnonblock _ _) _ _ |- _ ] => invert H
+ | [ H : stmnt_runp _ _ _ (Vblock _ _) _ _ |- _ ] => invert H
+ | [ H : stmnt_runp _ _ _ Vskip _ _ |- _ ] => invert H
+ | [ H : stmnt_runp _ _ _ (Vcond _ _ _) _ _ |- _ ] => invert H
+ | [ H : stmnt_runp _ _ _ (Vcase _ (_ :: _) _) _ _ |- _ ] => invert H
+ | [ H : stmnt_runp _ _ _ (Vcase _ [] _) _ _ |- _ ] => invert H
+
+ | [ H1 : expr_runp _ ?asr ?asa ?e _,
+ H2 : expr_runp _ ?asr ?asa ?e _ |- _ ] =>
+ learn (expr_runp_determinate H1 H2)
+
+ | [ H1 : location_is _ ?asr ?asa ?e _,
+ H2 : location_is _ ?asr ?asa ?e _ |- _ ] =>
+ learn (location_is_determinate H1 H2)
+
+ | [ H1 : forall asr1 asa1, stmnt_runp _ ?asr0 ?asa0 ?s asr1 asa1 -> _,
+ H2 : stmnt_runp _ ?asr0 ?asa0 ?s _ _ |- _ ] =>
+ learn (H1 _ _ H2)
+ end; crush).
+Qed.
+Hint Resolve stmnt_runp_determinate : verilog.
+
+Lemma mi_stepp_determinate :
+ forall f asr0 asa0 m asr1 asa1,
+ mi_stepp f asr0 asa0 m asr1 asa1 ->
+ forall asr1' asa1',
+ mi_stepp f asr0 asa0 m asr1' asa1' ->
+ asr1' = asr1 /\ asa1' = asa1.
+Proof.
+ intros. destruct m; invert H; invert H0;
+
+ repeat (try match goal with
+ | [ H1 : stmnt_runp _ ?asr0 ?asa0 ?s _ _,
+ H2 : stmnt_runp _ ?asr0 ?asa0 ?s _ _ |- _ ] =>
+ learn (stmnt_runp_determinate H1 H2)
+ end; crush).
+Qed.
+
+Lemma mis_stepp_determinate :
+ forall f asr0 asa0 m asr1 asa1,
+ mis_stepp f asr0 asa0 m asr1 asa1 ->
+ forall asr1' asa1',
+ mis_stepp f asr0 asa0 m asr1' asa1' ->
+ asr1' = asr1 /\ asa1' = asa1.
+Proof.
+ induction 1; intros;
+
+ repeat (try match goal with
+ | [ H : mis_stepp _ _ _ [] _ _ |- _ ] => invert H
+ | [ H : mis_stepp _ _ _ ( _ :: _ ) _ _ |- _ ] => invert H
+
+ | [ H1 : mi_stepp _ ?asr0 ?asa0 ?s _ _,
+ H2 : mi_stepp _ ?asr0 ?asa0 ?s _ _ |- _ ] =>
+ learn (mi_stepp_determinate H1 H2)
+
+ | [ H1 : forall asr1 asa1, mis_stepp _ ?asr0 ?asa0 ?m asr1 asa1 -> _,
+ H2 : mis_stepp _ ?asr0 ?asa0 ?m _ _ |- _ ] =>
+ learn (H1 _ _ H2)
+ end; crush).
+Qed.
+
+Lemma semantics_determinate :
+ forall (p: program), Smallstep.determinate (semantics p).
+Proof.
+ intros. constructor; set (ge := Globalenvs.Genv.globalenv p); simplify.
+ - invert H; invert H0; constructor. (* Traces are always empty *)
+ - invert H; invert H0; crush. assert (f = f0) by (destruct f; destruct f0; auto); subst.
+ pose proof (mis_stepp_determinate H5 H15).
+ crush.
+ - constructor. invert H; crush.
+ - invert H; invert H0; unfold ge0, ge1 in *; crush.
+ - red; simplify; intro; invert H0; invert H; crush.
+ - invert H; invert H0; crush.
+Qed.
diff --git a/test/array.c b/test/array.c
index e33d47b..7d78a61 100644
--- a/test/array.c
+++ b/test/array.c
@@ -1,4 +1,7 @@
int main() {
- int x[5] = {1, 2, 3, 4, 5};
- return x[2];
+ int x[3] = {1, 2, 3};
+ int sum = 0, incr = 1;
+ for (int i = 0; i < 3; i=i+incr)
+ sum += x[i];
+ return sum;
}
diff --git a/test/matrix.c b/test/matrix.c
index 2bb17e7..daa00ae 100644
--- a/test/matrix.c
+++ b/test/matrix.c
@@ -1,10 +1,8 @@
-#define N 4
-
-void matrix_multiply(int first[N][N], int second[N][N], int multiply[N][N]) {
+void matrix_multiply(int first[2][2], int second[2][2], int multiply[2][2]) {
int sum = 0;
- for (int c = 0; c < N; c++) {
- for (int d = 0; d < N; d++) {
- for (int k = 0; k < N; k++) {
+ for (int c = 0; c < 2; c++) {
+ for (int d = 0; d < 2; d++) {
+ for (int k = 0; k < 2; k++) {
sum = sum + first[c][k]*second[k][d];
}
multiply[c][d] = sum;
@@ -14,9 +12,9 @@ void matrix_multiply(int first[N][N], int second[N][N], int multiply[N][N]) {
}
int main() {
- int f[N][N] = {{1, 2, 3, 4}, {1, 2, 3, 4}, {1, 2, 3, 4}, {1, 2, 3, 4}};
- int s[N][N] = {{5, 6, 7, 8}, {5, 6, 7, 8}, {5, 6, 7, 8}, {5, 6, 7, 8}};
- int m[N][N] = {{0, 0, 0, 0}, {0, 0, 0, 0}, {0, 0, 0, 0}, {0, 0, 0, 0}};
+ int f[2][2] = {{1, 2}, {3, 4}};
+ int s[2][2] = {{5, 6}, {7, 8}};
+ int m[2][2] = {{0, 0}, {0, 0}};
matrix_multiply(f, s, m);
return m[1][1];