Skip to content

Commit 335e57a

Browse files
use vm_compute instead of Ltac to check instruction bounds
1 parent 8293653 commit 335e57a

File tree

4 files changed

+9
-8
lines changed

4 files changed

+9
-8
lines changed

compiler/src/compiler/CompilerInvariant.v

+4-1
Original file line numberDiff line numberDiff line change
@@ -2,6 +2,7 @@ Require Import coqutil.Tactics.rewr.
22
Require Import coqutil.Map.Interface coqutil.Map.Properties.
33
Require Import coqutil.Word.Interface coqutil.Word.Properties.
44
Require Import coqutil.Byte.
5+
Require Import riscv.Utility.bverify.
56
Require Import bedrock2.Array.
67
Require Import bedrock2.Map.SeparationLogic.
78
Require Import compiler.SeparationLogic.
@@ -174,7 +175,7 @@ Section Pipeline1.
174175
required_stack_space <= word.unsigned (word.sub (stack_pastend ml) (stack_start ml)) / bytes_per_word /\
175176
word.unsigned ml.(code_start) + Z.of_nat (List.length (instrencode instrs)) <=
176177
word.unsigned ml.(code_pastend) /\
177-
Forall (fun i => verify i iset \/ valid_InvalidInstruction i) instrs /\
178+
bvalidInstructions iset instrs = true /\
178179
(imem ml.(code_start) ml.(code_pastend) instrs *
179180
mem_available ml.(heap_start) ml.(heap_pastend) *
180181
mem_available ml.(stack_start) ml.(stack_pastend))%sep initial.(getMem) /\
@@ -211,6 +212,7 @@ Section Pipeline1.
211212
destruct mlOk.
212213
destruct M0 as [v M0].
213214
* apply ptsto_bytes_to_program; try assumption.
215+
eapply bvalidInstructions_valid. assumption.
214216
* unfold ptsto_bytes in Imem.
215217
eapply ptsto_bytes_range; try eassumption.
216218
+ unfold imem in *.
@@ -220,6 +222,7 @@ Section Pipeline1.
220222
eapply iff1ToEq.
221223
destruct mlOk.
222224
eapply ptsto_bytes_to_program; try eassumption.
225+
eapply bvalidInstructions_valid. assumption.
223226
- eapply @ll_inv_is_invariant; eassumption.
224227
- eapply ll_inv_implies_prefix_of_good. eassumption.
225228
Qed.

deps/riscv-coq

end2end/src/end2end/End2EndLightbulb.v

+2-5
Original file line numberDiff line numberDiff line change
@@ -289,9 +289,6 @@ Proof.
289289

290290
- vm_compute. intros. discriminate.
291291
- vm_compute. intros. discriminate.
292-
- (* Here we prove that all > 700 instructions are valid, using Ltac.
293-
If this becomes a bottleneck, we'll have to do this in Gallina in the compile function. *)
294-
unfold lightbulb_insts. repeat (apply Forall_cons || apply Forall_nil).
295-
all: left; vm_compute; try intuition discriminate.
296292
- vm_compute. reflexivity.
297-
Time Qed. (* takes more than 25s *)
293+
- vm_compute. reflexivity.
294+
Time Qed. (* takes ca 2s *)

end2end/src/end2end/End2EndPipeline.v

+2-1
Original file line numberDiff line numberDiff line change
@@ -15,6 +15,7 @@ Require riscv.Platform.Memory.
1515
Require Import riscv.Spec.PseudoInstructions.
1616
Require Import riscv.Proofs.EncodeBound.
1717
Require Import riscv.Proofs.DecodeEncode.
18+
Require Import riscv.Utility.bverify.
1819
Require Import riscv.Platform.Run.
1920
Require Import riscv.Utility.MkMachineWidth.
2021
Require Import riscv.Utility.Monads. Import MonadNotations.
@@ -307,7 +308,7 @@ Section Connect.
307308
required_stack_space <= word.unsigned (word.sub (stack_pastend ml) (stack_start ml)) / bytes_per_word ->
308309
word.unsigned (code_start ml) + Z.of_nat (Datatypes.length (instrencode instrs)) <=
309310
word.unsigned (code_pastend ml) ->
310-
Forall (fun i : Instruction => verify i iset \/ valid_InvalidInstruction i) instrs ->
311+
bvalidInstructions iset instrs = true ->
311312
valid_src_funs funimplsList = true ->
312313
(* Assumptions on the Kami level: *)
313314
kami_mem_contains_bytes (instrencode instrs) ml.(code_start) memInit ->

0 commit comments

Comments
 (0)