|
| 1 | +Require Import Coq.ZArith.ZArith. |
| 2 | +Require Import riscv.Spec.Decode. |
| 3 | +Require Import riscv.Utility.Encode. |
| 4 | +Require Import riscv.Utility.Utility. |
| 5 | + |
| 6 | +(* same verification functions as in Encode, but computable: *) |
| 7 | + |
| 8 | +Require Import Coq.derive.Derive. |
| 9 | + |
| 10 | +Lemma andb_spec: forall P Q p q, |
| 11 | + Bool.reflect P p -> |
| 12 | + Bool.reflect Q q -> |
| 13 | + Bool.reflect (P /\ Q) (p && q). |
| 14 | +Proof. |
| 15 | + intros. destruct H; destruct H0; constructor; intuition idtac. |
| 16 | +Qed. |
| 17 | + |
| 18 | +Lemma orb_spec: forall P Q p q, |
| 19 | + Bool.reflect P p -> |
| 20 | + Bool.reflect Q q -> |
| 21 | + Bool.reflect (P \/ Q) (p || q). |
| 22 | +Proof. |
| 23 | + intros. destruct H; destruct H0; constructor; intuition idtac. |
| 24 | +Qed. |
| 25 | + |
| 26 | +Lemma negb_spec: forall P p, |
| 27 | + Bool.reflect P p -> |
| 28 | + Bool.reflect (~ P) (negb p). |
| 29 | +Proof. |
| 30 | + intros. destruct H; constructor; intuition idtac. |
| 31 | +Qed. |
| 32 | + |
| 33 | +Lemma true_spec: Bool.reflect True true. Proof. constructor. constructor. Qed. |
| 34 | + |
| 35 | +Lemma false_spec: Bool.reflect False false. Proof. constructor. auto. Qed. |
| 36 | + |
| 37 | +Local Ltac t := |
| 38 | + intros; |
| 39 | + repeat match goal with |
| 40 | + | x := _ |- _ => subst x |
| 41 | + end; |
| 42 | + repeat first |
| 43 | + [ eapply andb_spec |
| 44 | + | eapply orb_spec |
| 45 | + | eapply Z.leb_spec0 |
| 46 | + | eapply Z.ltb_spec0 |
| 47 | + | eapply Z.eqb_spec |
| 48 | + | eapply true_spec |
| 49 | + | eapply false_spec ]. |
| 50 | + |
| 51 | +Derive bverify_Invalid SuchThat |
| 52 | + (forall i: Z, Bool.reflect (verify_Invalid i) (bverify_Invalid i)) |
| 53 | +As bverify_Invalid_spec. |
| 54 | + subst bverify_Invalid. |
| 55 | + intros. unfold verify_Invalid. eapply Bool.ReflectF. auto. |
| 56 | +Defined. |
| 57 | + |
| 58 | +Derive bverify_R SuchThat |
| 59 | + (forall (opcode: MachineInt)(rd rs1 rs2: Register)(funct3 funct7: MachineInt), |
| 60 | + Bool.reflect (verify_R opcode rd rs1 rs2 funct3 funct7) |
| 61 | + (bverify_R opcode rd rs1 rs2 funct3 funct7)) |
| 62 | +As bverify_R_spec. t. Defined. |
| 63 | + |
| 64 | +Derive bverify_R_atomic SuchThat |
| 65 | + (forall (opcode: MachineInt)(rd rs1 rs2: Register)(funct3 aqrl funct5: MachineInt), |
| 66 | + Bool.reflect (verify_R_atomic opcode rd rs1 rs2 funct3 aqrl funct5) |
| 67 | + (bverify_R_atomic opcode rd rs1 rs2 funct3 aqrl funct5)) |
| 68 | +As bverify_R_atomic_spec. t. Defined. |
| 69 | + |
| 70 | +Derive bverify_I SuchThat |
| 71 | + (forall (opcode: MachineInt)(rd rs1: Register)(funct3: MachineInt)(oimm12: Z), |
| 72 | + Bool.reflect (verify_I opcode rd rs1 funct3 oimm12) |
| 73 | + (bverify_I opcode rd rs1 funct3 oimm12)) |
| 74 | +As bverify_I_spec. t. Defined. |
| 75 | + |
| 76 | +Derive bverify_I_shift_57 SuchThat |
| 77 | + (forall (opcode: MachineInt)(rd rs1: Register)(shamt5 funct3 funct7: MachineInt), |
| 78 | + Bool.reflect (verify_I_shift_57 opcode rd rs1 shamt5 funct3 funct7) |
| 79 | + (bverify_I_shift_57 opcode rd rs1 shamt5 funct3 funct7)) |
| 80 | +As bverify_I_shift_57_spec. t. Defined. |
| 81 | + |
| 82 | +Derive bverify_I_shift_66 SuchThat |
| 83 | + (forall (bitwidth: Z)(opcode: MachineInt)(rd rs1: Register)(shamt6 funct3 funct6: MachineInt), |
| 84 | + Bool.reflect (verify_I_shift_66 bitwidth opcode rd rs1 shamt6 funct3 funct6) |
| 85 | + (bverify_I_shift_66 bitwidth opcode rd rs1 shamt6 funct3 funct6)) |
| 86 | +As bverify_I_shift_66_spec. t. Defined. |
| 87 | + |
| 88 | +Derive bverify_I_system SuchThat |
| 89 | + (forall (opcode: MachineInt)(rd rs1: Register)(funct3 funct12: MachineInt), |
| 90 | + Bool.reflect (verify_I_system opcode rd rs1 funct3 funct12) |
| 91 | + (bverify_I_system opcode rd rs1 funct3 funct12)) |
| 92 | +As bverify_I_system_spec. t. Defined. |
| 93 | + |
| 94 | +Derive bverify_S SuchThat |
| 95 | + (forall (opcode: MachineInt)(rs1 rs2: Register)(funct3: MachineInt)(simm12: Z), |
| 96 | + Bool.reflect (verify_S opcode rs1 rs2 funct3 simm12) |
| 97 | + (bverify_S opcode rs1 rs2 funct3 simm12)) |
| 98 | +As bverify_S_spec. t. Defined. |
| 99 | + |
| 100 | +Derive bverify_SB SuchThat |
| 101 | + (forall (opcode: MachineInt)(rs1 rs2: Register)(funct3: MachineInt)(sbimm12: Z), |
| 102 | + Bool.reflect (verify_SB opcode rs1 rs2 funct3 sbimm12) |
| 103 | + (bverify_SB opcode rs1 rs2 funct3 sbimm12)) |
| 104 | +As bverify_SB_spec. t. Defined. |
| 105 | + |
| 106 | +Derive bverify_U SuchThat |
| 107 | + (forall (opcode: MachineInt)(rd: Register)(imm20: Z), |
| 108 | + Bool.reflect (verify_U opcode rd imm20) |
| 109 | + (bverify_U opcode rd imm20)) |
| 110 | +As bverify_U_spec. t. Defined. |
| 111 | + |
| 112 | +Derive bverify_UJ SuchThat |
| 113 | + (forall (opcode: MachineInt)(rd: Register)(jimm20: Z), |
| 114 | + Bool.reflect (verify_UJ opcode rd jimm20) |
| 115 | + (bverify_UJ opcode rd jimm20)) |
| 116 | +As bverify_UJ_spec. t. Defined. |
| 117 | + |
| 118 | +Derive bverify_Fence SuchThat |
| 119 | + (forall (opcode: MachineInt)(rd rs1: Register)(funct3 prd scc msb4: MachineInt), |
| 120 | + Bool.reflect (verify_Fence opcode rd rs1 funct3 prd scc msb4) |
| 121 | + (bverify_Fence opcode rd rs1 funct3 prd scc msb4)) |
| 122 | +As bverify_Fence_spec. t. Defined. |
| 123 | + |
| 124 | +Derive bverify_FenceI SuchThat |
| 125 | + (forall (opcode: MachineInt)(rd rs1: Register)(funct3 imm12: MachineInt), |
| 126 | + Bool.reflect (verify_FenceI opcode rd rs1 funct3 imm12) |
| 127 | + (bverify_FenceI opcode rd rs1 funct3 imm12)) |
| 128 | +As bverify_FenceI_spec. t. Defined. |
| 129 | + |
| 130 | +Definition bVerifier(bitwidth: Z): InstructionMapper bool := {| |
| 131 | + map_Invalid := bverify_Invalid; |
| 132 | + map_R := bverify_R; |
| 133 | + map_R_atomic := bverify_R_atomic; |
| 134 | + map_I := bverify_I; |
| 135 | + map_I_shift_57 := bverify_I_shift_57; |
| 136 | + map_I_shift_66 := bverify_I_shift_66 bitwidth; |
| 137 | + map_I_system := bverify_I_system; |
| 138 | + map_S := bverify_S; |
| 139 | + map_SB := bverify_SB; |
| 140 | + map_U := bverify_U; |
| 141 | + map_UJ := bverify_UJ; |
| 142 | + map_Fence := bverify_Fence; |
| 143 | + map_FenceI := bverify_FenceI; |
| 144 | +|}. |
| 145 | + |
| 146 | +Definition brespects_bounds(bitwidth: Z): Instruction -> bool := |
| 147 | + apply_InstructionMapper (bVerifier bitwidth). |
| 148 | + |
| 149 | +Lemma brespects_bounds_spec: forall bitwidth i, |
| 150 | + Bool.reflect (respects_bounds bitwidth i) (brespects_bounds bitwidth i). |
| 151 | +Proof. |
| 152 | + intros. unfold respects_bounds, brespects_bounds. |
| 153 | + destruct i as [i | i | i | i | i | i | i | i | i | i ]; destruct i; simpl; |
| 154 | + eauto using bverify_Invalid_spec, |
| 155 | + bverify_R_spec, |
| 156 | + bverify_R_atomic_spec, |
| 157 | + bverify_I_spec, |
| 158 | + bverify_I_shift_57_spec, |
| 159 | + bverify_I_shift_66_spec, |
| 160 | + bverify_I_system_spec, |
| 161 | + bverify_S_spec, |
| 162 | + bverify_SB_spec, |
| 163 | + bverify_U_spec, |
| 164 | + bverify_UJ_spec, |
| 165 | + bverify_Fence_spec, |
| 166 | + bverify_FenceI_spec. |
| 167 | +Qed. |
| 168 | + |
| 169 | +Derive iset_eqb SuchThat |
| 170 | + (forall iset1 iset2: InstructionSet, Bool.reflect (iset1 = iset2) (iset_eqb iset1 iset2)) |
| 171 | +As iset_eqb_spec. |
| 172 | + Unshelve. 2: { intros iset1 iset2. destruct iset1; destruct iset2; shelve. } |
| 173 | + subst iset_eqb. |
| 174 | + destruct iset1; destruct iset2; |
| 175 | + first [ eapply Bool.ReflectT; reflexivity |
| 176 | + | eapply Bool.ReflectF; intro C; discriminate C ]. |
| 177 | +Defined. |
| 178 | + |
| 179 | +Derive bverify_iset SuchThat |
| 180 | + (forall (inst: Instruction)(iset: InstructionSet), |
| 181 | + Bool.reflect (verify_iset inst iset) (bverify_iset inst iset)) |
| 182 | +As bverify_iset_spec. |
| 183 | + intros. |
| 184 | + let f := open_constr:(ltac:(intro inst'; destruct inst'; shelve): |
| 185 | + Instruction -> InstructionSet -> bool) in |
| 186 | + unify bverify_iset f. |
| 187 | + destruct inst; simpl; t; eapply iset_eqb_spec. |
| 188 | +Defined. |
| 189 | + |
| 190 | +Definition bverify(inst: Instruction)(iset: InstructionSet): bool := |
| 191 | + andb (brespects_bounds (bitwidth iset) inst) (bverify_iset inst iset). |
| 192 | + |
| 193 | +Lemma bverify_spec: forall inst iset, Bool.reflect (verify inst iset) (bverify inst iset). |
| 194 | +Proof. t; [eapply brespects_bounds_spec | eapply bverify_iset_spec]. Qed. |
| 195 | + |
| 196 | +Definition valid_InvalidInstruction(instr: Decode.Instruction): Prop := |
| 197 | + exists n, (0 <= n < 2 ^ 32)%Z /\ instr = InvalidInstruction n. |
| 198 | + |
| 199 | +Definition bvalid_InvalidInstruction(instr: Decode.Instruction): bool := |
| 200 | + match instr with |
| 201 | + | InvalidInstruction n => andb (0 <=? n)%Z (n <? 2 ^ 32)%Z |
| 202 | + | _ => false |
| 203 | + end. |
| 204 | + |
| 205 | +Lemma bvalid_InvalidInstruction_spec: forall instr, |
| 206 | + Bool.reflect (valid_InvalidInstruction instr) (bvalid_InvalidInstruction instr). |
| 207 | +Proof. |
| 208 | + destruct instr; cbn -[Z.pow]; cycle -1. |
| 209 | + 1: destruct (0 <=? inst)%Z eqn: E1; [|simpl]. |
| 210 | + 1: destruct (inst <? 2 ^ 32)%Z eqn: E2; simpl. |
| 211 | + all: unfold valid_InvalidInstruction. |
| 212 | + { eapply Bool.ReflectT. eapply Z.leb_le in E1. eapply Z.ltb_lt in E2. eauto. } |
| 213 | + all: eapply Bool.ReflectF; intros (n & (A1 & A2) & B); inversion B; subst; clear B. |
| 214 | + { eapply Z.ltb_lt in A2. congruence. } |
| 215 | + { eapply Z.leb_le in A1. congruence. } |
| 216 | +Qed. |
| 217 | + |
| 218 | +Definition validInstruction(iset: InstructionSet)(i: Instruction): Prop := |
| 219 | + verify i iset \/ valid_InvalidInstruction i. |
| 220 | + |
| 221 | +Definition bvalidInstruction(iset: InstructionSet)(i: Instruction): bool := |
| 222 | + orb (bverify i iset) (bvalid_InvalidInstruction i). |
| 223 | + |
| 224 | +Lemma bvalidInstruction_spec: forall iset i, |
| 225 | + Bool.reflect (validInstruction iset i) (bvalidInstruction iset i). |
| 226 | +Proof. |
| 227 | + intros. eapply orb_spec. 1: eapply bverify_spec. 1: eapply bvalid_InvalidInstruction_spec. |
| 228 | +Qed. |
| 229 | + |
| 230 | +Definition validInstructions(iset: InstructionSet): list Instruction -> Prop := |
| 231 | + List.Forall (validInstruction iset). |
| 232 | + |
| 233 | +Definition bvalidInstructions(iset: InstructionSet): list Instruction -> bool := |
| 234 | + List.forallb (bvalidInstruction iset). |
| 235 | + |
| 236 | +Lemma bvalidInstructions_spec: forall iset instrs, |
| 237 | + Bool.reflect (validInstructions iset instrs) (bvalidInstructions iset instrs). |
| 238 | +Proof. |
| 239 | + intros. destruct (bvalidInstructions iset instrs) eqn: E; |
| 240 | + unfold bvalidInstructions, validInstructions in *. |
| 241 | + - eapply Bool.ReflectT. |
| 242 | + eapply List.Forall_forall. |
| 243 | + intros. |
| 244 | + eapply List.forallb_forall in E. 2: eassumption. |
| 245 | + eapply Bool.reflect_iff. 1: eapply bvalidInstruction_spec. assumption. |
| 246 | + - eapply Bool.ReflectF. |
| 247 | + intro C. |
| 248 | + pose proof (proj1 (List.Forall_forall _ _) C) as D. |
| 249 | + assert (forall i, List.In i instrs -> bvalidInstruction iset i = true) as F. { |
| 250 | + intros. eapply (proj1 (Bool.reflect_iff _ _ (bvalidInstruction_spec iset i))). |
| 251 | + eapply D. assumption. |
| 252 | + } |
| 253 | + eapply List.forallb_forall in F. |
| 254 | + congruence. |
| 255 | +Qed. |
| 256 | + |
| 257 | +Lemma bvalidInstructions_valid: forall iset instrs, |
| 258 | + bvalidInstructions iset instrs = true -> validInstructions iset instrs. |
| 259 | +Proof. intros. eapply Bool.reflect_iff. 1: eapply bvalidInstructions_spec. assumption. Qed. |
0 commit comments