From a454294cdf4f674a1bcb0f08408d64f9c1d99bbe Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Fri, 11 Aug 2023 00:48:03 -0500 Subject: [PATCH 01/42] defined assembly-level leakage trace, defined leakage_of_instr --- src/riscv/Spec/Decode.v | 355 +++++++++++++++++++++++++++++++++++++++- 1 file changed, 354 insertions(+), 1 deletion(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 4bbf55e..066db1a 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -55,6 +55,25 @@ Inductive InstructionM64 : Type := | Remuw (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM64 | InvalidM64 : InstructionM64. +(* I believe division is often not implemented in a constant-time way, so this will have to + be changed at some point. *) +Inductive LeakageM64 : Type := +| Mulw_leakage +| Divw_leakage +| Divuw_leakage +| Remw_leakage +| Remuw_leakage. + +Definition leakage_of_instr_M64 (instr : InstructionM64) : option LeakageM64 := + match instr with + | Mulw _ _ _ => Some Mulw_leakage + | Divw _ _ _ => Some Divw_leakage + | Divuw _ _ _ => Some Divw_leakage + | Remw _ _ _ => Some Remw_leakage + | Remuw _ _ _ => Some Remuw_leakage + | InvalidM64 => None + end. + Inductive InstructionM : Type := | Mul (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM | Mulh (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM @@ -66,6 +85,30 @@ Inductive InstructionM : Type := | Remu (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM | InvalidM : InstructionM. +(* This will also have to change, due to division beingn implemented in a not-constant-time way. *) +Inductive LeakageM : Type := +| Mul_leakage +| Mulh_leakage +| Mulhsu_leakage +| Mulhu_leakage +| Div_leakage +| Divu_leakage +| Rem_leakage +| Remu_leakage. + +Definition leakage_of_instr_M (instr : InstructionM) : option LeakageM := + match instr with + | Mul _ _ _ => Some Mul_leakage + | Mulh _ _ _ => Some Mulh_leakage + | Mulhsu _ _ _ => Some Mulhsu_leakage + | Mulhu _ _ _ => Some Mulhu_leakage + | Div _ _ _ => Some Div_leakage + | Divu _ _ _ => Some Divu_leakage + | Rem _ _ _ => Some Rem_leakage + | Remu _ _ _ => Some Remu_leakage + | InvalidM => None + end. + Inductive InstructionI64 : Type := | Ld (rd : Register) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) : InstructionI64 @@ -85,6 +128,78 @@ Inductive InstructionI64 : Type := | Sraw (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionI64 | InvalidI64 : InstructionI64. +Inductive LeakageI64 : Type := +| Ld_leakage (addr: Utility.Utility.MachineInt) +| Lwu_leakage (addr: Utility.Utility.MachineInt) +| Addiw_leakage +| Slliw_leakage +| Srliw_leakage +| Sraiw_leakage +| Sd_leakage (addr: Utility.Utility.MachineInt) +| Addw_leakage +| Subw_leakage +| Sllw_leakage +| Srlw_leakage +| Sraw_leakage. Search Utility.Utility.MachineInt. + +Definition leakage_of_instr_I64 (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionI64) : option LeakageI64 := + match instr with + | Ld _ rs1 oimm12 => Some (Ld_leakage (getRegister rs1 + oimm12)) + | Lwu _ rs1 oimm12 => Some (Lwu_leakage (getRegister rs1 + oimm12)) + | Addiw _ _ _ => Some Addiw_leakage + | Slliw _ _ _ => Some Slliw_leakage + | Srliw _ _ _ => Some Srliw_leakage + | Sraiw _ _ _ => Some Sraiw_leakage + | Sd rs1 _ simm12 => Some (Sd_leakage (getRegister rs1 + simm12)) + | Addw _ _ _ => Some Addw_leakage + | Subw _ _ _ => Some Subw_leakage + | Sllw _ _ _ => Some Sllw_leakage + | Srlw _ _ _ => Some Srlw_leakage + | Sraw _ _ _ => Some Sraw_leakage + | InvalidI64 => None + end. + +Inductive LeakageI : Type := +| Lb_leakage (addr: Utility.Utility.MachineInt) +| Lh_leakage (addr: Utility.Utility.MachineInt) +| Lw_leakage (addr: Utility.Utility.MachineInt) +| Lbu_leakage (addr: Utility.Utility.MachineInt) +| Lhu_leakage (addr: Utility.Utility.MachineInt) +| Fence_leakage (* unsure about this one. *) +| Fence_i_leakage +| Addi_leakage +| Slli_leakage +| Slti_leakage +| Sltiu_leakage +| Xori_leakage +| Ori_leakage +| Andi_leakage +| Srli_leakage +| Srai_leakage +| Auipc_leakage +| Sb_leakage (addr: Utility.Utility.MachineInt) +| Sh_leakage (addr: Utility.Utility.MachineInt) +| Sw_leakage (addr: Utility.Utility.MachineInt) +| Add_leakage +| Sub_leakage +| Sll_leakage +| Slt_leakage +| Sltu_leakage +| Xor_leakage +| Srl_leakage +| Sra_leakage +| Or_leakage +| And_leakage +| Lui_leakage +| Beq_leakage (branch: bool) (* unsure whether this should be here - i think that andres said that having this argument would make my life harder *) +| Bne_leakage (branch: bool) +| Blt_leakage (branch: bool) +| Bge_leakage (branch: bool) +| Bltu_leakage (branch: bool) +| Bgeu_leakage (branch: bool) +| Jalr_leakage (* unsure whether i should add the location we're jumping to here - or on the branches. i think not? *) +| Jal_leakage. + Inductive InstructionI : Type := | Lb (rd : Register) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) : InstructionI @@ -147,7 +262,52 @@ Inductive InstructionI : Type := | Jalr (rd : Register) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) : InstructionI | Jal (rd : Register) (jimm20 : Utility.Utility.MachineInt) : InstructionI - | InvalidI : InstructionI. + | InvalidI : InstructionI. Search Z. Compute (Utility.regToZ_unsigned 0). + +(* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) +Definition leakage_of_instr_I (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionI) : option LeakageI := + match instr with + | Lb _ rs1 oimm12 => Some (Lb_leakage (getRegister rs1 + oimm12)) + | Lh _ rs1 oimm12 => Some (Lh_leakage (getRegister rs1 + oimm12)) + | Lw _ rs1 oimm12 => Some (Lw_leakage (getRegister rs1 + oimm12)) + | Lbu _ rs1 oimm12 => Some (Lbu_leakage (getRegister rs1 + oimm12)) + | Lhu _ rs1 oimm12 => Some (Lhu_leakage (getRegister rs1 + oimm12)) + | Fence _ _ => Some Fence_leakage + | Fence_i => Some Fence_i_leakage + | Addi _ _ _ => Some Addi_leakage + | Slli _ _ _ => Some Slli_leakage + | Slti _ _ _ => Some Slti_leakage + | Sltiu _ _ _ => Some Sltiu_leakage + | Xori _ _ _ => Some Xori_leakage + | Ori _ _ _ => Some Ori_leakage + | Andi _ _ _ => Some Andi_leakage + | Srli _ _ _ => Some Srli_leakage + | Srai _ _ _ => Some Srai_leakage + | Auipc _ _ => Some Auipc_leakage + | Sb rs1 _ simm12 => Some (Sb_leakage (getRegister rs1 + simm12)) + | Sh rs1 _ simm12 => Some (Sh_leakage (getRegister rs1 + simm12)) + | Sw rs1 _ simm12 => Some (Sw_leakage (getRegister rs1 + simm12)) + | Add _ _ _ => Some Add_leakage + | Sub _ _ _ => Some Sub_leakage + | Sll _ _ _ => Some Sll_leakage + | Slt _ _ _ => Some Slt_leakage + | Sltu _ _ _ => Some Sltu_leakage + | Xor _ _ _ => Some Xor_leakage + | Srl _ _ _ => Some Srl_leakage + | Sra _ _ _ => Some Sra_leakage + | Or _ _ _ => Some Or_leakage + | And _ _ _ => Some And_leakage + | Lui _ _ => Some Lui_leakage + | Beq rs1 rs2 _ => Some (Beq_leakage (getRegister rs1 =? getRegister rs2)) + | Bne rs1 rs2 _ => Some (Bne_leakage (negb (getRegister rs1 =? getRegister rs2))) + | Blt rs1 rs2 _ => Some (Blt_leakage (getRegister rs1 Some (Bge_leakage (getRegister rs1 >=? getRegister rs2)) + | Bltu rs1 rs2 _ => Some (Bltu_leakage (getRegister rs1 <=? getRegister rs2)) + | Bgeu rs1 rs2 _ => Some (Bgeu_leakage (getRegister rs1 >=? getRegister rs2)) + | Jalr _ _ _ => Some Jalr_leakage + | Jal _ _ => Some Jal_leakage + | InvalidI => None + end. Inductive InstructionF64 : Type := | Fcvt_l_s (rd : Register) (rs1 : FPRegister) (rm : RoundMode) : InstructionF64 @@ -156,6 +316,49 @@ Inductive InstructionF64 : Type := | Fcvt_s_lu (rd : FPRegister) (rs1 : Register) (rm : RoundMode) : InstructionF64 | InvalidF64 : InstructionF64. +Inductive LeakageF64 : Type := +| Fcvt_l_s_leakage +| Fcvt_lu_s_leakage +| Fcvt_s_l_leakage +| Fcvt_s_lu_leakage. + +Definition leakage_of_instr_F64 (instr : InstructionF64) : option LeakageF64 := + match instr with + | Fcvt_l_s _ _ _ => Some Fcvt_l_s_leakage + | Fcvt_lu_s _ _ _ => Some Fcvt_lu_s_leakage + | Fcvt_s_l _ _ _ => Some Fcvt_s_l_leakage + | Fcvt_s_lu _ _ _ => Some Fcvt_s_lu_leakage + | InvalidF64 => None + end. + +Inductive LeakageF : Type := +| Flw_leakage (addr: Utility.Utility.MachineInt) +| Fsw_leakage (addr: Utility.Utility.MachineInt) +| Fmadd_s_leakage +| Fmsub_s_leakage +| Fnmsub_s_leakage +| Fnmadd_s_leakage +| Fadd_s_leakage +| Fsub_s_leakage +| Fmul_s_leakage +| Fdiv_s_leakage +| Fsqrt_s_leakage +| Fsgnj_s_leakage +| Fsgnjn_s_leakage +| Fsgnjx_s_leakage +| Fmin_s_leakage +| Fmax_s_leakage +| Fcvt_w_s_leakage +| Fcvt_wu_s_leakage +| Fmv_x_w_leakage +| Feq_s_leakage +| Flt_s_leakage +| Fle_s_leakage +| Fclass_s_leakage +| Fcvt_s_w_leakage +| Fcvt_s_wu_leakage +| Fmv_w_x_leakage. + Inductive InstructionF : Type := | Flw (rd : FPRegister) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) : InstructionF @@ -205,6 +408,54 @@ Inductive InstructionF : Type := | Fmv_w_x (rd : FPRegister) (rs1 : Register) : InstructionF | InvalidF : InstructionF. +Definition leakage_of_instr_F (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionF) : option LeakageF := + match instr with + | Flw _ rs1 oimm12 => Some (Flw_leakage (getRegister rs1 + oimm12)) + | Fsw rs1 _ simm12 => Some (Fsw_leakage (getRegister rs1 + simm12)) + | Fmadd_s _ _ _ _ _ => Some Fmadd_s_leakage + | Fmsub_s _ _ _ _ _ => Some Fmsub_s_leakage + | Fnmsub_s _ _ _ _ _ => Some Fnmsub_s_leakage + | Fnmadd_s _ _ _ _ _ => Some Fnmadd_s_leakage + | Fadd_s _ _ _ _ => Some Fadd_s_leakage + | Fsub_s _ _ _ _ => Some Fsub_s_leakage + | Fmul_s _ _ _ _ => Some Fmul_s_leakage + | Fdiv_s _ _ _ _ => Some Fdiv_s_leakage + | Fsqrt_s _ _ _ => Some Fsqrt_s_leakage + | Fsgnj_s _ _ _ => Some Fsgnj_s_leakage + | Fsgnjn_s _ _ _ => Some Fsgnjn_s_leakage + | Fsgnjx_s _ _ _ => Some Fsgnjx_s_leakage + | Fmin_s _ _ _ => Some Fmin_s_leakage + | Fmax_s _ _ _ => Some Fmax_s_leakage + | Fcvt_w_s _ _ _ => Some Fcvt_w_s_leakage + | Fcvt_wu_s _ _ _ => Some Fcvt_wu_s_leakage + | Fmv_x_w _ _ => Some Fmv_x_w_leakage + | Feq_s _ _ _ => Some Feq_s_leakage + | Flt_s _ _ _ => Some Flt_s_leakage + | Fle_s _ _ _ => Some Fle_s_leakage + | Fclass_s _ _ => Some Fclass_s_leakage + | Fcvt_s_w _ _ _ => Some Fcvt_s_w_leakage + | Fcvt_s_wu _ _ _ => Some Fcvt_s_wu_leakage + | Fmv_w_x _ _ => Some Fmv_w_x_leakage + | InvalidF => None + end. + +(* unsure what some of these mean, so unsure whether I have the right definition here. *) +Inductive LeakageCSR : Type := +| Ecall_leakage +| Ebreak_leakage +| Uret_leakage +| Sret_leakage +| Mret_leakage +| Wfi_leakage +| Sfence_vma_leakage +| Csrrw_leakage +| Csrrs_leakage +| Csrrc_leakage +| Csrrwi_leakage +| Csrrsi_leakage +| Csrrci_leakage. + +(* what are these, and why is it hard to find documentation for them? *) Inductive InstructionCSR : Type := | Ecall : InstructionCSR | Ebreak : InstructionCSR @@ -230,6 +481,38 @@ Inductive InstructionCSR : Type := : InstructionCSR | InvalidCSR : InstructionCSR. +Definition leakage_of_instr_CSR (instr : InstructionCSR) : option LeakageCSR := + match instr with + | Ecall => Some Ecall_leakage + | Ebreak => Some Ebreak_leakage + | Uret => Some Uret_leakage + | Sret => Some Uret_leakage + | Mret => Some Mret_leakage + | Wfi => Some Wfi_leakage + | Sfence_vma _ _ => Some Sfence_vma_leakage + | Csrrw _ _ _ => Some Csrrw_leakage + | Csrrs _ _ _ => Some Csrrs_leakage + | Csrrc _ _ _ => Some Csrrc_leakage + | Csrrwi _ _ _ => Some Csrrwi_leakage + | Csrrsi _ _ _ => Some Csrrsi_leakage + | Csrrci _ _ _ => Some Csrrci_leakage + | InvalidCSR => None + end. + +(* do we care about aqrl here? *) +Inductive LeakageA64 : Type := +| Lr_d_leakage (addr : Utility.Utility.MachineInt) +| Sc_d_leakage (addr : Utility.Utility.MachineInt) (* behavior of this depends on whether there is a reservation on addr... *) +| Amoswap_d_leakage (addr : Utility.Utility.MachineInt) +| Amoadd_d_leakage (addr : Utility.Utility.MachineInt) +| Amoand_d_leakage (addr : Utility.Utility.MachineInt) +| Amoor_d_leakage (addr : Utility.Utility.MachineInt) +| Amoxor_d_leakage (addr : Utility.Utility.MachineInt) +| Amomax_d_leakage (addr : Utility.Utility.MachineInt) +| Amomaxu_d_leakage (addr : Utility.Utility.MachineInt) +| Amomin_d_leakage (addr : Utility.Utility.MachineInt) +| Amominu_d_leakage (addr : Utility.Utility.MachineInt). + Inductive InstructionA64 : Type := | Lr_d (rd : Register) (rs1 : Register) (aqrl : Utility.Utility.MachineInt) : InstructionA64 @@ -265,6 +548,35 @@ Inductive InstructionA64 : Type := : InstructionA64 | InvalidA64 : InstructionA64. +Definition leakage_of_instr_A64 (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionA64) : option LeakageA64 := + match instr with + | Lr_d _ rs1 _ => Some (Lr_d_leakage (getRegister rs1)) + | Sc_d _ rs1 _ _ => Some (Sc_d_leakage (getRegister rs1)) + | Amoswap_d _ rs1 _ _ => Some (Amoswap_d_leakage (getRegister rs1)) + | Amoadd_d _ rs1 _ _ => Some (Amoadd_d_leakage (getRegister rs1)) + | Amoand_d _ rs1 _ _ => Some (Amoand_d_leakage (getRegister rs1)) + | Amoor_d _ rs1 _ _ => Some (Amoor_d_leakage (getRegister rs1)) + | Amoxor_d _ rs1 _ _ => Some (Amoxor_d_leakage (getRegister rs1)) + | Amomax_d _ rs1 _ _ => Some (Amomax_d_leakage (getRegister rs1)) + | Amomaxu_d _ rs1 _ _ => Some (Amomaxu_d_leakage (getRegister rs1)) + | Amomin_d _ rs1 _ _ => Some (Amomin_d_leakage (getRegister rs1)) + | Amominu_d _ rs1 _ _ => Some (Amominu_d_leakage (getRegister rs1)) + | InvalidA64 => None + end. + +Inductive LeakageA : Type := +| Lr_w_leakage (addr : Utility.Utility.MachineInt) +| Sc_w_leakage (addr : Utility.Utility.MachineInt) +| Amoswap_w_leakage (addr : Utility.Utility.MachineInt) +| Amoadd_w_leakage (addr : Utility.Utility.MachineInt) +| Amoand_w_leakage (addr : Utility.Utility.MachineInt) +| Amoor_w_leakage (addr : Utility.Utility.MachineInt) +| Amoxor_w_leakage (addr : Utility.Utility.MachineInt) +| Amomax_w_leakage (addr : Utility.Utility.MachineInt) +| Amomaxu_w_leakage (addr : Utility.Utility.MachineInt) +| Amomin_w_leakage (addr : Utility.Utility.MachineInt) +| Amominu_w_leakage (addr : Utility.Utility.MachineInt). + Inductive InstructionA : Type := | Lr_w (rd : Register) (rs1 : Register) (aqrl : Utility.Utility.MachineInt) : InstructionA @@ -300,6 +612,22 @@ Inductive InstructionA : Type := : InstructionA | InvalidA : InstructionA. +Definition leakage_of_instr_A (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionA) : option LeakageA := + match instr with + | Lr_w _ rs1 _ => Some (Lr_w_leakage (getRegister rs1)) + | Sc_w _ rs1 _ _ => Some (Sc_w_leakage (getRegister rs1)) + | Amoswap_w _ rs1 _ _ => Some (Amoswap_w_leakage (getRegister rs1)) + | Amoadd_w _ rs1 _ _ => Some (Amoadd_w_leakage (getRegister rs1)) + | Amoand_w _ rs1 _ _ => Some (Amoand_w_leakage (getRegister rs1)) + | Amoor_w _ rs1 _ _ => Some (Amoor_w_leakage (getRegister rs1)) + | Amoxor_w _ rs1 _ _ => Some (Amoxor_w_leakage (getRegister rs1)) + | Amomax_w _ rs1 _ _ => Some (Amomax_w_leakage (getRegister rs1)) + | Amomaxu_w _ rs1 _ _ => Some (Amomaxu_w_leakage (getRegister rs1)) + | Amomin_w _ rs1 _ _ => Some (Amomin_w_leakage (getRegister rs1)) + | Amominu_w _ rs1 _ _ => Some (Amominu_w_leakage (getRegister rs1)) + | InvalidA => None + end. + Inductive Instruction : Type := | IInstruction (iInstruction : InstructionI) : Instruction | MInstruction (mInstruction : InstructionM) : Instruction @@ -312,6 +640,31 @@ Inductive Instruction : Type := | CSRInstruction (csrInstruction : InstructionCSR) : Instruction | InvalidInstruction (inst : Utility.Utility.MachineInt) : Instruction. +Inductive LeakageEvent : Type := +| ILeakage (iLeakage : LeakageI) +| MLeakage (mLeakage : LeakageM) +| ALeakage (aLeakage : LeakageA) +| FLeakage (fLeakage : LeakageF) +| I64Leakage (i64Leakage : LeakageI64) +| M64Leakage (m64Leakage : LeakageM64) +| A64Leakage (a64Leakage : LeakageA64) +| F64Leakage (f64Leakage : LeakageF64) +| CSRLeakage (csrLeakage : LeakageCSR). + +Definition leakage_of_instr (getRegister : Register -> Utility.Utility.MachineInt) (instr : Instruction) : option LeakageEvent := + match instr with + | IInstruction instr => option_map ILeakage (leakage_of_instr_I getRegister instr) + | MInstruction instr => option_map MLeakage (leakage_of_instr_M instr) + | AInstruction instr => option_map ALeakage (leakage_of_instr_A getRegister instr) + | FInstruction instr => option_map FLeakage (leakage_of_instr_F getRegister instr) + | I64Instruction instr => option_map I64Leakage (leakage_of_instr_I64 getRegister instr) + | M64Instruction instr => option_map M64Leakage (leakage_of_instr_M64 instr) + | A64Instruction instr => option_map A64Leakage (leakage_of_instr_A64 getRegister instr) + | F64Instruction instr => option_map F64Leakage (leakage_of_instr_F64 instr) + | CSRInstruction instr => option_map CSRLeakage (leakage_of_instr_CSR instr) + | InvalidInstruction _ => None + end. + (* Converted value declarations: *) (* Skipping instance `Spec.Decode.Eq___InstructionSet' of class From dfb34e2f8183c6eac3a35f8ac501e683091609cb Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Fri, 25 Aug 2023 19:59:41 -0500 Subject: [PATCH 02/42] updated leakage_of_instr to use a monad --- src/riscv/Spec/Decode.v | 356 +++++++++++++++++++++------------------- 1 file changed, 187 insertions(+), 169 deletions(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 066db1a..bd82263 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -27,6 +27,12 @@ Require Coq.Lists.List. Require Import Coq.ZArith.BinInt. Require Utility.Utility. +Require Import riscv.Utility.Monads. +Require Import riscv.Utility.MonadNotations. +Section WithMonad. + Context {M : Type -> Type} {MM : Monad M}. + Context (getRegister : Register -> M Utility.Utility.MachineInt). + (* Converted type declarations: *) Inductive InstructionSet : Type := @@ -62,16 +68,17 @@ Inductive LeakageM64 : Type := | Divw_leakage | Divuw_leakage | Remw_leakage -| Remuw_leakage. +| Remuw_leakage +| InvalidM64_leakage. -Definition leakage_of_instr_M64 (instr : InstructionM64) : option LeakageM64 := +Definition leakage_of_instr_M64 (instr : InstructionM64) : M LeakageM64 := match instr with - | Mulw _ _ _ => Some Mulw_leakage - | Divw _ _ _ => Some Divw_leakage - | Divuw _ _ _ => Some Divw_leakage - | Remw _ _ _ => Some Remw_leakage - | Remuw _ _ _ => Some Remuw_leakage - | InvalidM64 => None + | Mulw _ _ _ => Return Mulw_leakage + | Divw _ _ _ => Return Divw_leakage + | Divuw _ _ _ => Return Divw_leakage + | Remw _ _ _ => Return Remw_leakage + | Remuw _ _ _ => Return Remuw_leakage + | InvalidM64 => Return InvalidM64_leakage end. Inductive InstructionM : Type := @@ -94,19 +101,20 @@ Inductive LeakageM : Type := | Div_leakage | Divu_leakage | Rem_leakage -| Remu_leakage. +| Remu_leakage +| InvalidM_leakage. -Definition leakage_of_instr_M (instr : InstructionM) : option LeakageM := +Definition leakage_of_instr_M (instr : InstructionM) : M LeakageM := match instr with - | Mul _ _ _ => Some Mul_leakage - | Mulh _ _ _ => Some Mulh_leakage - | Mulhsu _ _ _ => Some Mulhsu_leakage - | Mulhu _ _ _ => Some Mulhu_leakage - | Div _ _ _ => Some Div_leakage - | Divu _ _ _ => Some Divu_leakage - | Rem _ _ _ => Some Rem_leakage - | Remu _ _ _ => Some Remu_leakage - | InvalidM => None + | Mul _ _ _ => Return Mul_leakage + | Mulh _ _ _ => Return Mulh_leakage + | Mulhsu _ _ _ => Return Mulhsu_leakage + | Mulhu _ _ _ => Return Mulhu_leakage + | Div _ _ _ => Return Div_leakage + | Divu _ _ _ => Return Divu_leakage + | Rem _ _ _ => Return Rem_leakage + | Remu _ _ _ => Return Remu_leakage + | InvalidM => Return InvalidM_leakage end. Inductive InstructionI64 : Type := @@ -140,23 +148,24 @@ Inductive LeakageI64 : Type := | Subw_leakage | Sllw_leakage | Srlw_leakage -| Sraw_leakage. Search Utility.Utility.MachineInt. +| Sraw_leakage +| InvalidI64_leakage. Locate "<- ; _ ". Check Bind. -Definition leakage_of_instr_I64 (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionI64) : option LeakageI64 := +Definition leakage_of_instr_I64 (instr : InstructionI64) : M LeakageI64 := match instr with - | Ld _ rs1 oimm12 => Some (Ld_leakage (getRegister rs1 + oimm12)) - | Lwu _ rs1 oimm12 => Some (Lwu_leakage (getRegister rs1 + oimm12)) - | Addiw _ _ _ => Some Addiw_leakage - | Slliw _ _ _ => Some Slliw_leakage - | Srliw _ _ _ => Some Srliw_leakage - | Sraiw _ _ _ => Some Sraiw_leakage - | Sd rs1 _ simm12 => Some (Sd_leakage (getRegister rs1 + simm12)) - | Addw _ _ _ => Some Addw_leakage - | Subw _ _ _ => Some Subw_leakage - | Sllw _ _ _ => Some Sllw_leakage - | Srlw _ _ _ => Some Srlw_leakage - | Sraw _ _ _ => Some Sraw_leakage - | InvalidI64 => None + | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (rs1_val + oimm12))) + | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (rs1_val + oimm12))) + | Addiw _ _ _ => Return Addiw_leakage + | Slliw _ _ _ => Return Slliw_leakage + | Srliw _ _ _ => Return Srliw_leakage + | Sraiw _ _ _ => Return Sraiw_leakage + | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (rs1_val + simm12))) + | Addw _ _ _ => Return Addw_leakage + | Subw _ _ _ => Return Subw_leakage + | Sllw _ _ _ => Return Sllw_leakage + | Srlw _ _ _ => Return Srlw_leakage + | Sraw _ _ _ => Return Sraw_leakage + | InvalidI64 => Return InvalidI64_leakage end. Inductive LeakageI : Type := @@ -198,7 +207,8 @@ Inductive LeakageI : Type := | Bltu_leakage (branch: bool) | Bgeu_leakage (branch: bool) | Jalr_leakage (* unsure whether i should add the location we're jumping to here - or on the branches. i think not? *) -| Jal_leakage. +| Jal_leakage +| InvalidI_leakage. Inductive InstructionI : Type := | Lb (rd : Register) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) @@ -265,48 +275,48 @@ Inductive InstructionI : Type := | InvalidI : InstructionI. Search Z. Compute (Utility.regToZ_unsigned 0). (* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) -Definition leakage_of_instr_I (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionI) : option LeakageI := +Definition leakage_of_instr_I (instr : InstructionI) : M LeakageI := match instr with - | Lb _ rs1 oimm12 => Some (Lb_leakage (getRegister rs1 + oimm12)) - | Lh _ rs1 oimm12 => Some (Lh_leakage (getRegister rs1 + oimm12)) - | Lw _ rs1 oimm12 => Some (Lw_leakage (getRegister rs1 + oimm12)) - | Lbu _ rs1 oimm12 => Some (Lbu_leakage (getRegister rs1 + oimm12)) - | Lhu _ rs1 oimm12 => Some (Lhu_leakage (getRegister rs1 + oimm12)) - | Fence _ _ => Some Fence_leakage - | Fence_i => Some Fence_i_leakage - | Addi _ _ _ => Some Addi_leakage - | Slli _ _ _ => Some Slli_leakage - | Slti _ _ _ => Some Slti_leakage - | Sltiu _ _ _ => Some Sltiu_leakage - | Xori _ _ _ => Some Xori_leakage - | Ori _ _ _ => Some Ori_leakage - | Andi _ _ _ => Some Andi_leakage - | Srli _ _ _ => Some Srli_leakage - | Srai _ _ _ => Some Srai_leakage - | Auipc _ _ => Some Auipc_leakage - | Sb rs1 _ simm12 => Some (Sb_leakage (getRegister rs1 + simm12)) - | Sh rs1 _ simm12 => Some (Sh_leakage (getRegister rs1 + simm12)) - | Sw rs1 _ simm12 => Some (Sw_leakage (getRegister rs1 + simm12)) - | Add _ _ _ => Some Add_leakage - | Sub _ _ _ => Some Sub_leakage - | Sll _ _ _ => Some Sll_leakage - | Slt _ _ _ => Some Slt_leakage - | Sltu _ _ _ => Some Sltu_leakage - | Xor _ _ _ => Some Xor_leakage - | Srl _ _ _ => Some Srl_leakage - | Sra _ _ _ => Some Sra_leakage - | Or _ _ _ => Some Or_leakage - | And _ _ _ => Some And_leakage - | Lui _ _ => Some Lui_leakage - | Beq rs1 rs2 _ => Some (Beq_leakage (getRegister rs1 =? getRegister rs2)) - | Bne rs1 rs2 _ => Some (Bne_leakage (negb (getRegister rs1 =? getRegister rs2))) - | Blt rs1 rs2 _ => Some (Blt_leakage (getRegister rs1 Some (Bge_leakage (getRegister rs1 >=? getRegister rs2)) - | Bltu rs1 rs2 _ => Some (Bltu_leakage (getRegister rs1 <=? getRegister rs2)) - | Bgeu rs1 rs2 _ => Some (Bgeu_leakage (getRegister rs1 >=? getRegister rs2)) - | Jalr _ _ _ => Some Jalr_leakage - | Jal _ _ => Some Jal_leakage - | InvalidI => None + | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (rs1_val + oimm12)) + | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (rs1_val + oimm12)) + | Lw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lw_leakage (rs1_val + oimm12)) + | Lbu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lbu_leakage (rs1_val + oimm12)) + | Lhu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lhu_leakage (rs1_val + oimm12)) + | Fence _ _ => Return Fence_leakage + | Fence_i => Return Fence_i_leakage + | Addi _ _ _ => Return Addi_leakage + | Slli _ _ _ => Return Slli_leakage + | Slti _ _ _ => Return Slti_leakage + | Sltiu _ _ _ => Return Sltiu_leakage + | Xori _ _ _ => Return Xori_leakage + | Ori _ _ _ => Return Ori_leakage + | Andi _ _ _ => Return Andi_leakage + | Srli _ _ _ => Return Srli_leakage + | Srai _ _ _ => Return Srai_leakage + | Auipc _ _ => Return Auipc_leakage + | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (rs1_val + simm12)) + | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (rs1_val + simm12)) + | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (rs1_val + simm12)) + | Add _ _ _ => Return Add_leakage + | Sub _ _ _ => Return Sub_leakage + | Sll _ _ _ => Return Sll_leakage + | Slt _ _ _ => Return Slt_leakage + | Sltu _ _ _ => Return Sltu_leakage + | Xor _ _ _ => Return Xor_leakage + | Srl _ _ _ => Return Srl_leakage + | Sra _ _ _ => Return Sra_leakage + | Or _ _ _ => Return Or_leakage + | And _ _ _ => Return And_leakage + | Lui _ _ => Return Lui_leakage + | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Beq_leakage (rs1_val =? rs2_val)) + | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bne_leakage (negb (rs1_val =? rs2_val))) + | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Blt_leakage (rs1_val rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bge_leakage (rs1_val >=? rs2_val)) + | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bltu_leakage (rs1_val <=? rs2_val)) + | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bgeu_leakage (rs1_val >=? rs2_val)) + | Jalr _ _ _ => Return Jalr_leakage + | Jal _ _ => Return Jal_leakage + | InvalidI => Return InvalidI_leakage end. Inductive InstructionF64 : Type := @@ -320,15 +330,16 @@ Inductive LeakageF64 : Type := | Fcvt_l_s_leakage | Fcvt_lu_s_leakage | Fcvt_s_l_leakage -| Fcvt_s_lu_leakage. +| Fcvt_s_lu_leakage +| InvalidF64_leakage. -Definition leakage_of_instr_F64 (instr : InstructionF64) : option LeakageF64 := +Definition leakage_of_instr_F64 (instr : InstructionF64) : M LeakageF64 := match instr with - | Fcvt_l_s _ _ _ => Some Fcvt_l_s_leakage - | Fcvt_lu_s _ _ _ => Some Fcvt_lu_s_leakage - | Fcvt_s_l _ _ _ => Some Fcvt_s_l_leakage - | Fcvt_s_lu _ _ _ => Some Fcvt_s_lu_leakage - | InvalidF64 => None + | Fcvt_l_s _ _ _ => Return Fcvt_l_s_leakage + | Fcvt_lu_s _ _ _ => Return Fcvt_lu_s_leakage + | Fcvt_s_l _ _ _ => Return Fcvt_s_l_leakage + | Fcvt_s_lu _ _ _ => Return Fcvt_s_lu_leakage + | InvalidF64 => Return InvalidF64_leakage end. Inductive LeakageF : Type := @@ -357,7 +368,8 @@ Inductive LeakageF : Type := | Fclass_s_leakage | Fcvt_s_w_leakage | Fcvt_s_wu_leakage -| Fmv_w_x_leakage. +| Fmv_w_x_leakage +| InvalidF_leakage. Inductive InstructionF : Type := | Flw (rd : FPRegister) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) @@ -408,35 +420,35 @@ Inductive InstructionF : Type := | Fmv_w_x (rd : FPRegister) (rs1 : Register) : InstructionF | InvalidF : InstructionF. -Definition leakage_of_instr_F (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionF) : option LeakageF := +Definition leakage_of_instr_F (instr : InstructionF) : M LeakageF := match instr with - | Flw _ rs1 oimm12 => Some (Flw_leakage (getRegister rs1 + oimm12)) - | Fsw rs1 _ simm12 => Some (Fsw_leakage (getRegister rs1 + simm12)) - | Fmadd_s _ _ _ _ _ => Some Fmadd_s_leakage - | Fmsub_s _ _ _ _ _ => Some Fmsub_s_leakage - | Fnmsub_s _ _ _ _ _ => Some Fnmsub_s_leakage - | Fnmadd_s _ _ _ _ _ => Some Fnmadd_s_leakage - | Fadd_s _ _ _ _ => Some Fadd_s_leakage - | Fsub_s _ _ _ _ => Some Fsub_s_leakage - | Fmul_s _ _ _ _ => Some Fmul_s_leakage - | Fdiv_s _ _ _ _ => Some Fdiv_s_leakage - | Fsqrt_s _ _ _ => Some Fsqrt_s_leakage - | Fsgnj_s _ _ _ => Some Fsgnj_s_leakage - | Fsgnjn_s _ _ _ => Some Fsgnjn_s_leakage - | Fsgnjx_s _ _ _ => Some Fsgnjx_s_leakage - | Fmin_s _ _ _ => Some Fmin_s_leakage - | Fmax_s _ _ _ => Some Fmax_s_leakage - | Fcvt_w_s _ _ _ => Some Fcvt_w_s_leakage - | Fcvt_wu_s _ _ _ => Some Fcvt_wu_s_leakage - | Fmv_x_w _ _ => Some Fmv_x_w_leakage - | Feq_s _ _ _ => Some Feq_s_leakage - | Flt_s _ _ _ => Some Flt_s_leakage - | Fle_s _ _ _ => Some Fle_s_leakage - | Fclass_s _ _ => Some Fclass_s_leakage - | Fcvt_s_w _ _ _ => Some Fcvt_s_w_leakage - | Fcvt_s_wu _ _ _ => Some Fcvt_s_wu_leakage - | Fmv_w_x _ _ => Some Fmv_w_x_leakage - | InvalidF => None + | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (rs1_val + oimm12)) + | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (rs1_val + simm12)) + | Fmadd_s _ _ _ _ _ => Return Fmadd_s_leakage + | Fmsub_s _ _ _ _ _ => Return Fmsub_s_leakage + | Fnmsub_s _ _ _ _ _ => Return Fnmsub_s_leakage + | Fnmadd_s _ _ _ _ _ => Return Fnmadd_s_leakage + | Fadd_s _ _ _ _ => Return Fadd_s_leakage + | Fsub_s _ _ _ _ => Return Fsub_s_leakage + | Fmul_s _ _ _ _ => Return Fmul_s_leakage + | Fdiv_s _ _ _ _ => Return Fdiv_s_leakage + | Fsqrt_s _ _ _ => Return Fsqrt_s_leakage + | Fsgnj_s _ _ _ => Return Fsgnj_s_leakage + | Fsgnjn_s _ _ _ => Return Fsgnjn_s_leakage + | Fsgnjx_s _ _ _ => Return Fsgnjx_s_leakage + | Fmin_s _ _ _ => Return Fmin_s_leakage + | Fmax_s _ _ _ => Return Fmax_s_leakage + | Fcvt_w_s _ _ _ => Return Fcvt_w_s_leakage + | Fcvt_wu_s _ _ _ => Return Fcvt_wu_s_leakage + | Fmv_x_w _ _ => Return Fmv_x_w_leakage + | Feq_s _ _ _ => Return Feq_s_leakage + | Flt_s _ _ _ => Return Flt_s_leakage + | Fle_s _ _ _ => Return Fle_s_leakage + | Fclass_s _ _ => Return Fclass_s_leakage + | Fcvt_s_w _ _ _ => Return Fcvt_s_w_leakage + | Fcvt_s_wu _ _ _ => Return Fcvt_s_wu_leakage + | Fmv_w_x _ _ => Return Fmv_w_x_leakage + | InvalidF => Return InvalidF_leakage end. (* unsure what some of these mean, so unsure whether I have the right definition here. *) @@ -453,7 +465,8 @@ Inductive LeakageCSR : Type := | Csrrc_leakage | Csrrwi_leakage | Csrrsi_leakage -| Csrrci_leakage. +| Csrrci_leakage +| InvalidCSR_leakage. (* what are these, and why is it hard to find documentation for them? *) Inductive InstructionCSR : Type := @@ -481,22 +494,22 @@ Inductive InstructionCSR : Type := : InstructionCSR | InvalidCSR : InstructionCSR. -Definition leakage_of_instr_CSR (instr : InstructionCSR) : option LeakageCSR := +Definition leakage_of_instr_CSR (instr : InstructionCSR) : M LeakageCSR := match instr with - | Ecall => Some Ecall_leakage - | Ebreak => Some Ebreak_leakage - | Uret => Some Uret_leakage - | Sret => Some Uret_leakage - | Mret => Some Mret_leakage - | Wfi => Some Wfi_leakage - | Sfence_vma _ _ => Some Sfence_vma_leakage - | Csrrw _ _ _ => Some Csrrw_leakage - | Csrrs _ _ _ => Some Csrrs_leakage - | Csrrc _ _ _ => Some Csrrc_leakage - | Csrrwi _ _ _ => Some Csrrwi_leakage - | Csrrsi _ _ _ => Some Csrrsi_leakage - | Csrrci _ _ _ => Some Csrrci_leakage - | InvalidCSR => None + | Ecall => Return Ecall_leakage + | Ebreak => Return Ebreak_leakage + | Uret => Return Uret_leakage + | Sret => Return Uret_leakage + | Mret => Return Mret_leakage + | Wfi => Return Wfi_leakage + | Sfence_vma _ _ => Return Sfence_vma_leakage + | Csrrw _ _ _ => Return Csrrw_leakage + | Csrrs _ _ _ => Return Csrrs_leakage + | Csrrc _ _ _ => Return Csrrc_leakage + | Csrrwi _ _ _ => Return Csrrwi_leakage + | Csrrsi _ _ _ => Return Csrrsi_leakage + | Csrrci _ _ _ => Return Csrrci_leakage + | InvalidCSR => Return InvalidCSR_leakage end. (* do we care about aqrl here? *) @@ -511,7 +524,8 @@ Inductive LeakageA64 : Type := | Amomax_d_leakage (addr : Utility.Utility.MachineInt) | Amomaxu_d_leakage (addr : Utility.Utility.MachineInt) | Amomin_d_leakage (addr : Utility.Utility.MachineInt) -| Amominu_d_leakage (addr : Utility.Utility.MachineInt). +| Amominu_d_leakage (addr : Utility.Utility.MachineInt) +| InvalidA64_leakage. Inductive InstructionA64 : Type := | Lr_d (rd : Register) (rs1 : Register) (aqrl : Utility.Utility.MachineInt) @@ -548,20 +562,20 @@ Inductive InstructionA64 : Type := : InstructionA64 | InvalidA64 : InstructionA64. -Definition leakage_of_instr_A64 (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionA64) : option LeakageA64 := +Definition leakage_of_instr_A64 (instr : InstructionA64) : M LeakageA64 := match instr with - | Lr_d _ rs1 _ => Some (Lr_d_leakage (getRegister rs1)) - | Sc_d _ rs1 _ _ => Some (Sc_d_leakage (getRegister rs1)) - | Amoswap_d _ rs1 _ _ => Some (Amoswap_d_leakage (getRegister rs1)) - | Amoadd_d _ rs1 _ _ => Some (Amoadd_d_leakage (getRegister rs1)) - | Amoand_d _ rs1 _ _ => Some (Amoand_d_leakage (getRegister rs1)) - | Amoor_d _ rs1 _ _ => Some (Amoor_d_leakage (getRegister rs1)) - | Amoxor_d _ rs1 _ _ => Some (Amoxor_d_leakage (getRegister rs1)) - | Amomax_d _ rs1 _ _ => Some (Amomax_d_leakage (getRegister rs1)) - | Amomaxu_d _ rs1 _ _ => Some (Amomaxu_d_leakage (getRegister rs1)) - | Amomin_d _ rs1 _ _ => Some (Amomin_d_leakage (getRegister rs1)) - | Amominu_d _ rs1 _ _ => Some (Amominu_d_leakage (getRegister rs1)) - | InvalidA64 => None + | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage rs1_val) + | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage rs1_val) + | Amoswap_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_d_leakage rs1_val) + | Amoadd_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_d_leakage rs1_val) + | Amoand_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_d_leakage rs1_val) + | Amoor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_d_leakage rs1_val) + | Amoxor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_d_leakage rs1_val) + | Amomax_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_d_leakage rs1_val) + | Amomaxu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_d_leakage rs1_val) + | Amomin_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_d_leakage rs1_val) + | Amominu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_d_leakage rs1_val) + | InvalidA64 => Return InvalidA64_leakage end. Inductive LeakageA : Type := @@ -575,7 +589,8 @@ Inductive LeakageA : Type := | Amomax_w_leakage (addr : Utility.Utility.MachineInt) | Amomaxu_w_leakage (addr : Utility.Utility.MachineInt) | Amomin_w_leakage (addr : Utility.Utility.MachineInt) -| Amominu_w_leakage (addr : Utility.Utility.MachineInt). +| Amominu_w_leakage (addr : Utility.Utility.MachineInt) +| InvalidA_leakage. Inductive InstructionA : Type := | Lr_w (rd : Register) (rs1 : Register) (aqrl : Utility.Utility.MachineInt) @@ -612,20 +627,20 @@ Inductive InstructionA : Type := : InstructionA | InvalidA : InstructionA. -Definition leakage_of_instr_A (getRegister : Register -> Utility.Utility.MachineInt) (instr : InstructionA) : option LeakageA := +Definition leakage_of_instr_A (instr : InstructionA) : M LeakageA := match instr with - | Lr_w _ rs1 _ => Some (Lr_w_leakage (getRegister rs1)) - | Sc_w _ rs1 _ _ => Some (Sc_w_leakage (getRegister rs1)) - | Amoswap_w _ rs1 _ _ => Some (Amoswap_w_leakage (getRegister rs1)) - | Amoadd_w _ rs1 _ _ => Some (Amoadd_w_leakage (getRegister rs1)) - | Amoand_w _ rs1 _ _ => Some (Amoand_w_leakage (getRegister rs1)) - | Amoor_w _ rs1 _ _ => Some (Amoor_w_leakage (getRegister rs1)) - | Amoxor_w _ rs1 _ _ => Some (Amoxor_w_leakage (getRegister rs1)) - | Amomax_w _ rs1 _ _ => Some (Amomax_w_leakage (getRegister rs1)) - | Amomaxu_w _ rs1 _ _ => Some (Amomaxu_w_leakage (getRegister rs1)) - | Amomin_w _ rs1 _ _ => Some (Amomin_w_leakage (getRegister rs1)) - | Amominu_w _ rs1 _ _ => Some (Amominu_w_leakage (getRegister rs1)) - | InvalidA => None + | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage rs1_val) + | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage rs1_val) + | Amoswap_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_w_leakage rs1_val) + | Amoadd_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_w_leakage rs1_val) + | Amoand_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_w_leakage rs1_val) + | Amoor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_w_leakage rs1_val) + | Amoxor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_w_leakage rs1_val) + | Amomax_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_w_leakage rs1_val) + | Amomaxu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_w_leakage rs1_val) + | Amomin_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_w_leakage rs1_val) + | Amominu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_w_leakage rs1_val) + | InvalidA => Return InvalidA_leakage end. Inductive Instruction : Type := @@ -649,21 +664,24 @@ Inductive LeakageEvent : Type := | M64Leakage (m64Leakage : LeakageM64) | A64Leakage (a64Leakage : LeakageA64) | F64Leakage (f64Leakage : LeakageF64) -| CSRLeakage (csrLeakage : LeakageCSR). +| CSRLeakage (csrLeakage : LeakageCSR) +| InvalidLeakage. Check @Bind. -Definition leakage_of_instr (getRegister : Register -> Utility.Utility.MachineInt) (instr : Instruction) : option LeakageEvent := +Definition leakage_of_instr (instr : Instruction) : M LeakageEvent := match instr with - | IInstruction instr => option_map ILeakage (leakage_of_instr_I getRegister instr) - | MInstruction instr => option_map MLeakage (leakage_of_instr_M instr) - | AInstruction instr => option_map ALeakage (leakage_of_instr_A getRegister instr) - | FInstruction instr => option_map FLeakage (leakage_of_instr_F getRegister instr) - | I64Instruction instr => option_map I64Leakage (leakage_of_instr_I64 getRegister instr) - | M64Instruction instr => option_map M64Leakage (leakage_of_instr_M64 instr) - | A64Instruction instr => option_map A64Leakage (leakage_of_instr_A64 getRegister instr) - | F64Instruction instr => option_map F64Leakage (leakage_of_instr_F64 instr) - | CSRInstruction instr => option_map CSRLeakage (leakage_of_instr_CSR instr) - | InvalidInstruction _ => None - end. + | IInstruction instr => l <- leakage_of_instr_I instr; Return (ILeakage l) + | MInstruction instr => l <- leakage_of_instr_M instr; Return (MLeakage l) + | AInstruction instr => l <- leakage_of_instr_A instr; Return (ALeakage l) + | FInstruction instr => l <- leakage_of_instr_F instr; Return (FLeakage l) + | I64Instruction instr => l <- leakage_of_instr_I64 instr; Return (I64Leakage l) + | M64Instruction instr => l <- leakage_of_instr_M64 instr; Return (M64Leakage l) + | A64Instruction instr => l <- leakage_of_instr_A64 instr; Return (A64Leakage l) + | F64Instruction instr => l <- leakage_of_instr_F64 instr; Return (F64Leakage l) + | CSRInstruction instr => l <- leakage_of_instr_CSR instr; Return (CSRLeakage l) + | InvalidInstruction _ => Return InvalidLeakage +end. + +End WithMonad. (* Converted value declarations: *) From 1ccf9826af8dfd0ef033445fdd94005a021031d1 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Fri, 3 Nov 2023 01:20:40 -0400 Subject: [PATCH 03/42] defined RiscvProgramWithLeakage (haven't used it for anything yet); added leakage trace to RiscvMachine --- src/riscv/Platform/RiscvMachine.v | 41 +++++++++++++++++++------------ src/riscv/Spec/Machine.v | 5 ++++ 2 files changed, 30 insertions(+), 16 deletions(-) diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index 8142c0c..7f3691d 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -141,40 +141,49 @@ Section Machine. getMem: Mem; getXAddrs: XAddrs; getLog: list LogItem; + getTrace: list LeakageEvent; }. Definition withRegs: Registers -> RiscvMachine -> RiscvMachine := - fun regs2 '(mkRiscvMachine regs1 pc nextPC mem xAddrs log) => - mkRiscvMachine regs2 pc nextPC mem xAddrs log. + fun regs2 '(mkRiscvMachine regs1 pc nextPC mem xAddrs log trace) => + mkRiscvMachine regs2 pc nextPC mem xAddrs log trace. Definition withPc: word -> RiscvMachine -> RiscvMachine := - fun pc2 '(mkRiscvMachine regs pc1 nextPC mem xAddrs log) => - mkRiscvMachine regs pc2 nextPC mem xAddrs log. + fun pc2 '(mkRiscvMachine regs pc1 nextPC mem xAddrs log trace) => + mkRiscvMachine regs pc2 nextPC mem xAddrs log trace. Definition withNextPc: word -> RiscvMachine -> RiscvMachine := - fun nextPC2 '(mkRiscvMachine regs pc nextPC1 mem xAddrs log) => - mkRiscvMachine regs pc nextPC2 mem xAddrs log. + fun nextPC2 '(mkRiscvMachine regs pc nextPC1 mem xAddrs log trace) => + mkRiscvMachine regs pc nextPC2 mem xAddrs log trace. Definition withMem: Mem -> RiscvMachine -> RiscvMachine := - fun mem2 '(mkRiscvMachine regs pc nextPC mem1 xAddrs log) => - mkRiscvMachine regs pc nextPC mem2 xAddrs log. + fun mem2 '(mkRiscvMachine regs pc nextPC mem1 xAddrs log trace) => + mkRiscvMachine regs pc nextPC mem2 xAddrs log trace. Definition withXAddrs: XAddrs -> RiscvMachine -> RiscvMachine := - fun xAddrs2 '(mkRiscvMachine regs pc nextPC mem xAddrs1 log) => - mkRiscvMachine regs pc nextPC mem xAddrs2 log. + fun xAddrs2 '(mkRiscvMachine regs pc nextPC mem xAddrs1 log trace) => + mkRiscvMachine regs pc nextPC mem xAddrs2 log trace. Definition withLog: list LogItem -> RiscvMachine -> RiscvMachine := - fun log2 '(mkRiscvMachine regs pc nextPC mem xAddrs log1) => - mkRiscvMachine regs pc nextPC mem xAddrs log2. + fun log2 '(mkRiscvMachine regs pc nextPC mem xAddrs log1 trace) => + mkRiscvMachine regs pc nextPC mem xAddrs log2 trace. Definition withLogItem: LogItem -> RiscvMachine -> RiscvMachine := - fun item '(mkRiscvMachine regs pc nextPC mem xAddrs log) => - mkRiscvMachine regs pc nextPC mem xAddrs (item :: log). + fun item '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => + mkRiscvMachine regs pc nextPC mem xAddrs (item :: log) trace. Definition withLogItems: list LogItem -> RiscvMachine -> RiscvMachine := - fun items '(mkRiscvMachine regs pc nextPC mem xAddrs log) => - mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log). + fun items '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => + mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log) trace. + Definition withTraceItem: LeakageEvent -> RiscvMachine -> RiscvMachine := + fun event '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => + mkRiscvMachine regs pc nextPC mem xAddrs log (event :: trace). + + Definition withTraceItems: list LeakageEvent -> RiscvMachine -> RiscvMachine := + fun events '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => + mkRiscvMachine regs pc nextPC mem xAddrs log (events ++ trace). + Definition Z32s_to_bytes(l: list Z): list byte := List.flat_map (fun z => HList.tuple.to_list (LittleEndian.split 4 z)) l. diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index 0bf9eb0..42e8b20 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -66,6 +66,11 @@ Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { endCycleEarly: forall A, M A; }. +Class RiscvProgramWithLeakage{M}{t}{MM}{MWt}:= + mkRiscvProgramWithLeakage { + RVP :> @RiscvProgram M t MM MWt; + leakInstr : Instruction -> M unit; + }. Class RiscvMachine`{MP: RiscvProgram} := mkRiscvMachine { (* checks that addr is aligned, and translates the (possibly virtual) addr to a physical From dc093fbaad405091a749ec15c17da2249cd6e2e7 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Fri, 3 Nov 2023 01:31:02 -0400 Subject: [PATCH 04/42] added instruction logging to run1 --- src/riscv/Platform/Run.v | 3 ++- src/riscv/Spec/Machine.v | 2 +- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index 75afe9f..ec56241 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -15,13 +15,14 @@ Section Riscv. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVP: RiscvProgram M mword}. + Context {RVP: RiscvProgramWithLeakage}. Context {RVS: RiscvMachine M mword}. Definition run1(iset: InstructionSet): M unit := pc <- getPC; inst <- loadWord Fetch pc; + logInstr (decode iset (combine 4 inst));; execute (decode iset (combine 4 inst));; endCycleNormal. diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index 42e8b20..9048dd4 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -69,7 +69,7 @@ Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { Class RiscvProgramWithLeakage{M}{t}{MM}{MWt}:= mkRiscvProgramWithLeakage { RVP :> @RiscvProgram M t MM MWt; - leakInstr : Instruction -> M unit; + logInstr : Instruction -> M unit; }. Class RiscvMachine`{MP: RiscvProgram} := mkRiscvMachine { From c0baecf228fe7f7e2889f3c6199b17231f2bf166 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Fri, 3 Nov 2023 03:40:55 -0400 Subject: [PATCH 05/42] added instruction logging to MetricMinimalMMIO; other things don't compile, but that's ok --- src/riscv/Platform/MaterializeRiscvProgram.v | 2 ++ src/riscv/Platform/MetricMinimalMMIO.v | 36 +++++++++++++++++++- src/riscv/Platform/MetricSane.v | 3 +- src/riscv/Platform/MinimalMMIO.v | 18 +++++++++- src/riscv/Spec/MetricPrimitives.v | 22 ++++++------ 5 files changed, 68 insertions(+), 13 deletions(-) diff --git a/src/riscv/Platform/MaterializeRiscvProgram.v b/src/riscv/Platform/MaterializeRiscvProgram.v index b8a522f..386aa76 100644 --- a/src/riscv/Platform/MaterializeRiscvProgram.v +++ b/src/riscv/Platform/MaterializeRiscvProgram.v @@ -27,6 +27,7 @@ Section Riscv. | GetPrivMode | SetPrivMode (_ : PrivMode) | Fence (_ : MachineInt) (_ : MachineInt) + | LogInstr (_ : Instruction) | GetPC | SetPC (_ : word) | StartCycle @@ -52,6 +53,7 @@ Section Riscv. | GetPrivMode => PrivMode | SetPrivMode _ => unit | Fence _ _ => unit + | LogInstr _ => unit | GetPC => word | SetPC _ => unit | StartCycle => unit diff --git a/src/riscv/Platform/MetricMinimalMMIO.v b/src/riscv/Platform/MetricMinimalMMIO.v index ac0318a..8d31115 100644 --- a/src/riscv/Platform/MetricMinimalMMIO.v +++ b/src/riscv/Platform/MetricMinimalMMIO.v @@ -30,7 +30,41 @@ Section Riscv. (* note: ext_spec does not have access to the metrics *) Context {mmio_spec: MMIOSpec}. - Local Notation M := (free action result). + (*should i change this to (list LeakageEvent -> list LeakageEvent) * ... ? *) + Definition action : Type := (MetricLog -> MetricLog) * riscv_primitive. + Definition result (a : action) := primitive_result (snd a). + Local Notation M := (free action result). Print act. + + Global Instance IsRiscvMachine: RiscvProgram M word := {| + getRegister a := act (id, GetRegister a) ret; + setRegister a b := act (id, SetRegister a b) ret; + loadByte a b := act (addMetricLoads 1, LoadByte a b) ret; + loadHalf a b := act (addMetricLoads 1, LoadHalf a b) ret; + loadWord a b := act (addMetricLoads 1, LoadWord a b) ret; + loadDouble a b := act (addMetricLoads 1, LoadDouble a b) ret; + storeByte a b c := act (addMetricStores 1, StoreByte a b c) ret; + storeHalf a b c := act (addMetricStores 1, StoreHalf a b c) ret; + storeWord a b c := act (addMetricStores 1, StoreWord a b c) ret; + storeDouble a b c := act (addMetricStores 1, StoreDouble a b c) ret; + makeReservation a := act (id, MakeReservation a) ret; + clearReservation a := act (id, ClearReservation a) ret; + checkReservation a := act (id, CheckReservation a) ret; + getCSRField f := act (id, GetCSRField f) ret; + setCSRField f v := act (id, SetCSRField f v) ret; + getPrivMode := act (id, GetPrivMode) ret; + setPrivMode m := act (id, SetPrivMode m) ret; + fence a b := act (id, Fence a b) ret; + getPC := act (id, GetPC) ret; + setPC a := act (addMetricJumps 1, SetPC a) ret; + endCycleNormal := act (addMetricInstructions 1, EndCycleNormal) ret; + endCycleEarly A := act (addMetricInstructions 1, EndCycleEarly A) ret; + |}. + + Check addMetricInstructions. + + Global Instance IsRiscvMachineWithLeakage: RiscvProgramWithLeakage := + {| + logInstr a := act (id, LogInstr a) ret; |}. Definition interp_action a metmach post := interpret_action (snd a) (metmach.(getMachine)) (fun r mach => diff --git a/src/riscv/Platform/MetricSane.v b/src/riscv/Platform/MetricSane.v index 422cb0b..05d75e1 100644 --- a/src/riscv/Platform/MetricSane.v +++ b/src/riscv/Platform/MetricSane.v @@ -20,7 +20,7 @@ Section Sane. Context {mem: map.map word byte}. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. + Context {RVM: RiscvProgramWithLeakage}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {mcomp_sat_ok: mcomp_sat_spec PRParams}. @@ -165,6 +165,7 @@ Section Sane. unfold run1. intros. apply Bind_sane; [apply getPC_sane|intros]. apply Bind_sane; [apply loadWord_sane|intros]. + apply Bind_sane; [apply logInstr_sane|intros]. apply Bind_sane; [apply execute_sane|intros]. apply endCycleNormal_sane. Qed. diff --git a/src/riscv/Platform/MinimalMMIO.v b/src/riscv/Platform/MinimalMMIO.v index dad0916..3f74776 100644 --- a/src/riscv/Platform/MinimalMMIO.v +++ b/src/riscv/Platform/MinimalMMIO.v @@ -22,6 +22,20 @@ Require Import riscv.Platform.Sane. Local Open Scope Z_scope. Local Open Scope bool_scope. +Definition Mtriv (x : Type) := x. +Definition trivialBind (A B : Type) (x : Mtriv A) (f : A -> B) : Mtriv B := f x. +Definition trivialReturn (A : Type) (a : A) : Mtriv A := a. +Print Monad. +Lemma trivial_left_identity : forall (A B : Type) (a : A) (f : A -> Mtriv B), trivialBind A B (trivialReturn A a) f = f a. +Proof. trivial. Qed. +Lemma trivial_right_identity : forall (A : Type) (m : Mtriv A), trivialBind A A m (trivialReturn A) = m. +Proof. trivial. Qed. +Lemma trivial_associativity : forall (A B C : Type) (m : Mtriv A) (f : A -> Mtriv B) (g : B -> Mtriv C), trivialBind B C (trivialBind A B m f) g = trivialBind A C m (fun x : A => trivialBind B C (f x) g). +Proof. trivial. Qed. + +Definition trivialMonad : Monad Mtriv := + {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. + Class MMIOSpec{width: Z}{BW: Bitwidth width}{word: word width}{Mem : map.map word byte} := { (* should not say anything about alignment, just whether it's in the MMIO range *) isMMIOAddr: word -> Prop; @@ -89,7 +103,7 @@ Section Riscv. else word.of_Z 0. Definition setReg(reg: Z)(v: word)(regs: Registers): Registers := - if ((0 RiscvMachine -> Prop) -> (RiscvMachine -> Prop) -> Prop := @@ -114,6 +128,8 @@ Section Riscv. postF tt (withNextPc (word.add mach.(getPc) (word.of_Z 4)) mach) | EndCycleNormal => fun postF postA => postF tt (updatePc mach) | EndCycleEarly _ => fun postF postA => postA (updatePc mach) (* ignores postF containing the continuation *) + | LogInstr i => fun (postF : unit -> RiscvMachine -> Prop) postA => + postF tt (withTraceItem (@leakage_of_instr Mtriv trivialMonad (fun reg => word.unsigned (getReg mach.(getRegs) reg)) i) mach) | MakeReservation _ | ClearReservation _ | CheckReservation _ diff --git a/src/riscv/Spec/MetricPrimitives.v b/src/riscv/Spec/MetricPrimitives.v index 4b13e25..dacdfd5 100644 --- a/src/riscv/Spec/MetricPrimitives.v +++ b/src/riscv/Spec/MetricPrimitives.v @@ -20,8 +20,9 @@ Section MetricPrimitives. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgram M word}. - Context {RVS: @RiscvMachine M word _ _ RVM}. + Context {RVM: RiscvProgramWithLeakage}. Print RiscvMachine. Print RiscvProgramWithLeakage. + Set Printing All. Print RVM. + Context {RVS: @RiscvMachine M word _ _ RVM.(RVP)}. (* monadic computations used for specifying the behavior of RiscvMachines should be "sane" in the sense that we never step to the empty set (that's not absence of failure, since @@ -54,6 +55,7 @@ Section MetricPrimitives. getPrivMode_sane: mcomp_sane getPrivMode; setPrivMode_sane: forall m, mcomp_sane (setPrivMode m); fence_sane: forall a b, mcomp_sane (fence a b); + logInstr_sane: forall a, mcomp_sane (logInstr a); getPC_sane: mcomp_sane getPC; setPC_sane: forall newPc, mcomp_sane (setPC newPc); endCycleNormal_sane: mcomp_sane endCycleNormal; @@ -108,15 +110,15 @@ Section MetricPrimitives. x = Register0 /\ post tt initialL) -> mcomp_sat (setRegister x v) initialL post; - spec_loadByte: spec_load 1 (Machine.loadByte (RiscvProgram := RVM)) Memory.loadByte; - spec_loadHalf: spec_load 2 (Machine.loadHalf (RiscvProgram := RVM)) Memory.loadHalf; - spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM)) Memory.loadWord; - spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM)) Memory.loadDouble; + spec_loadByte: spec_load 1 (Machine.loadByte (RiscvProgram := RVM.(RVP))) Memory.loadByte; + spec_loadHalf: spec_load 2 (Machine.loadHalf (RiscvProgram := RVM.(RVP))) Memory.loadHalf; + spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM.(RVP))) Memory.loadWord; + spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM.(RVP))) Memory.loadDouble; - spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM)) Memory.storeByte; - spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM)) Memory.storeHalf; - spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM)) Memory.storeWord; - spec_storeDouble: spec_store 8 (Machine.storeDouble (RiscvProgram := RVM)) Memory.storeDouble; + spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM.(RVP))) Memory.storeByte; + spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM.(RVP))) Memory.storeHalf; + spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM.(RVP))) Memory.storeWord; + spec_storeDouble: spec_store 8 (Machine.storeDouble (RiscvProgram := RVM.(RVP))) Memory.storeDouble; spec_getPC: forall (initialL: MetricRiscvMachine) (post: word -> MetricRiscvMachine -> Prop), post initialL.(getPc) initialL -> From dacf7ce07f02fb6811a13bf2fb80df270a0ba05f Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 5 Nov 2023 15:26:20 -0500 Subject: [PATCH 06/42] started getting rid of "logInstr" primitive; we have "leakEvent" now. --- src/riscv/Platform/MetricRiscvMachine.v | 8 ++++++++ src/riscv/Platform/RiscvMachine.v | 4 ++-- src/riscv/Spec/Machine.v | 2 +- src/riscv/Spec/MetricPrimitives.v | 6 +++++- 4 files changed, 16 insertions(+), 4 deletions(-) diff --git a/src/riscv/Platform/MetricRiscvMachine.v b/src/riscv/Platform/MetricRiscvMachine.v index c83abfa..c1b17b7 100644 --- a/src/riscv/Platform/MetricRiscvMachine.v +++ b/src/riscv/Platform/MetricRiscvMachine.v @@ -58,6 +58,14 @@ Section Machine. fun items '(mkMetricRiscvMachine mach mc) => (mkMetricRiscvMachine (withLogItems items mach) mc). + Definition withLeakageEvent: LeakageEvent -> MetricRiscvMachine -> MetricRiscvMachine := + fun event '(mkMetricRiscvMachine mach mc) => + (mkMetricRiscvMachine (withLeakageEvent event mach) mc). + + Definition withLeakageEvents: list LeakageEvent -> MetricRiscvMachine -> MetricRiscvMachine := + fun events '(mkMetricRiscvMachine mach mc) => + (mkMetricRiscvMachine (withLeakageEvents events mach) mc). + Definition forgetMetrics(m: MetricRiscvMachine): RiscvMachine := m.(getMachine). Definition addMetrics(m: RiscvMachine)(mc: MetricLog): MetricRiscvMachine := mkMetricRiscvMachine m mc. diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index 7f3691d..da30e7e 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -176,11 +176,11 @@ Section Machine. fun items '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log) trace. - Definition withTraceItem: LeakageEvent -> RiscvMachine -> RiscvMachine := + Definition withLeakageEvent: LeakageEvent -> RiscvMachine -> RiscvMachine := fun event '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => mkRiscvMachine regs pc nextPC mem xAddrs log (event :: trace). - Definition withTraceItems: list LeakageEvent -> RiscvMachine -> RiscvMachine := + Definition withLeakageEvents: list LeakageEvent -> RiscvMachine -> RiscvMachine := fun events '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => mkRiscvMachine regs pc nextPC mem xAddrs log (events ++ trace). diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index 9048dd4..b814426 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -69,7 +69,7 @@ Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { Class RiscvProgramWithLeakage{M}{t}{MM}{MWt}:= mkRiscvProgramWithLeakage { RVP :> @RiscvProgram M t MM MWt; - logInstr : Instruction -> M unit; + leakEvent : LeakageEvent -> M unit; }. Class RiscvMachine`{MP: RiscvProgram} := mkRiscvMachine { diff --git a/src/riscv/Spec/MetricPrimitives.v b/src/riscv/Spec/MetricPrimitives.v index dacdfd5..29c212b 100644 --- a/src/riscv/Spec/MetricPrimitives.v +++ b/src/riscv/Spec/MetricPrimitives.v @@ -55,7 +55,7 @@ Section MetricPrimitives. getPrivMode_sane: mcomp_sane getPrivMode; setPrivMode_sane: forall m, mcomp_sane (setPrivMode m); fence_sane: forall a b, mcomp_sane (fence a b); - logInstr_sane: forall a, mcomp_sane (logInstr a); + leakInstr_sane: forall a, mcomp_sane (leakEvent a); getPC_sane: mcomp_sane getPC; setPC_sane: forall newPc, mcomp_sane (setPC newPc); endCycleNormal_sane: mcomp_sane endCycleNormal; @@ -115,6 +115,10 @@ Section MetricPrimitives. spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM.(RVP))) Memory.loadWord; spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM.(RVP))) Memory.loadDouble; + spec_leakEvent: forall (initialL: MetricRiscvMachine) (post: unit -> MetricRiscvMachine -> Prop) (e : LeakageEvent), + post tt (withLeakageEvent e initialL) -> + mcomp_sat (leakEvent e) initialL post; + spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM.(RVP))) Memory.storeByte; spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM.(RVP))) Memory.storeHalf; spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM.(RVP))) Memory.storeWord; From e889f0004d8154fbe3f48b71b8b49500b701486a Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 5 Nov 2023 16:19:18 -0500 Subject: [PATCH 07/42] changed leakage_of_instr to accept a getRegister that returns a machine word --- src/riscv/Spec/Decode.v | 91 +++++++++++++++++++++-------------------- 1 file changed, 47 insertions(+), 44 deletions(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index bd82263..8e96075 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -25,13 +25,14 @@ Notation Opcode := BinInt.Z (only parsing). Require Coq.Init.Datatypes. Require Coq.Lists.List. Require Import Coq.ZArith.BinInt. -Require Utility.Utility. +Require Import Utility.Utility. Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Section WithMonad. Context {M : Type -> Type} {MM : Monad M}. - Context (getRegister : Register -> M Utility.Utility.MachineInt). + Context {mword: Type} {MW: MachineWidth mword}. + Context (getRegister : Register -> M mword). (* Converted type declarations: *) @@ -149,17 +150,19 @@ Inductive LeakageI64 : Type := | Sllw_leakage | Srlw_leakage | Sraw_leakage -| InvalidI64_leakage. Locate "<- ; _ ". Check Bind. +| InvalidI64_leakage. Locate "<- ; _ ". + +Print MW. Check add. Check regToZ_unsigned. Definition leakage_of_instr_I64 (instr : InstructionI64) : M LeakageI64 := match instr with - | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (rs1_val + oimm12))) - | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (rs1_val + oimm12))) + | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (regToZ_unsigned rs1_val + oimm12))) + | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (regToZ_unsigned rs1_val + oimm12))) | Addiw _ _ _ => Return Addiw_leakage | Slliw _ _ _ => Return Slliw_leakage | Srliw _ _ _ => Return Srliw_leakage | Sraiw _ _ _ => Return Sraiw_leakage - | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (rs1_val + simm12))) + | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (regToZ_unsigned rs1_val + simm12))) | Addw _ _ _ => Return Addw_leakage | Subw _ _ _ => Return Subw_leakage | Sllw _ _ _ => Return Sllw_leakage @@ -277,11 +280,11 @@ Inductive InstructionI : Type := (* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) Definition leakage_of_instr_I (instr : InstructionI) : M LeakageI := match instr with - | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (rs1_val + oimm12)) - | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (rs1_val + oimm12)) - | Lw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lw_leakage (rs1_val + oimm12)) - | Lbu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lbu_leakage (rs1_val + oimm12)) - | Lhu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lhu_leakage (rs1_val + oimm12)) + | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (regToZ_unsigned rs1_val + oimm12)) + | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (regToZ_unsigned rs1_val + oimm12)) + | Lw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lw_leakage (regToZ_unsigned rs1_val + oimm12)) + | Lbu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lbu_leakage (regToZ_unsigned rs1_val + oimm12)) + | Lhu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lhu_leakage (regToZ_unsigned rs1_val + oimm12)) | Fence _ _ => Return Fence_leakage | Fence_i => Return Fence_i_leakage | Addi _ _ _ => Return Addi_leakage @@ -294,9 +297,9 @@ Definition leakage_of_instr_I (instr : InstructionI) : M LeakageI := | Srli _ _ _ => Return Srli_leakage | Srai _ _ _ => Return Srai_leakage | Auipc _ _ => Return Auipc_leakage - | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (rs1_val + simm12)) - | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (rs1_val + simm12)) - | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (rs1_val + simm12)) + | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (regToZ_unsigned rs1_val + simm12)) + | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (regToZ_unsigned rs1_val + simm12)) + | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (regToZ_unsigned rs1_val + simm12)) | Add _ _ _ => Return Add_leakage | Sub _ _ _ => Return Sub_leakage | Sll _ _ _ => Return Sll_leakage @@ -308,12 +311,12 @@ Definition leakage_of_instr_I (instr : InstructionI) : M LeakageI := | Or _ _ _ => Return Or_leakage | And _ _ _ => Return And_leakage | Lui _ _ => Return Lui_leakage - | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Beq_leakage (rs1_val =? rs2_val)) - | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bne_leakage (negb (rs1_val =? rs2_val))) - | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Blt_leakage (rs1_val rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bge_leakage (rs1_val >=? rs2_val)) - | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bltu_leakage (rs1_val <=? rs2_val)) - | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bgeu_leakage (rs1_val >=? rs2_val)) + | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Beq_leakage (regToZ_unsigned rs1_val =? regToZ_unsigned rs2_val)) + | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bne_leakage (negb (regToZ_unsigned rs1_val =? regToZ_unsigned rs2_val))) + | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Blt_leakage (regToZ_signed rs1_val rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bge_leakage (regToZ_signed rs1_val >=? regToZ_signed rs2_val)) + | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bltu_leakage (regToZ_unsigned rs1_val <=? regToZ_unsigned rs2_val)) + | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bgeu_leakage (regToZ_unsigned rs1_val >=? regToZ_unsigned rs2_val)) | Jalr _ _ _ => Return Jalr_leakage | Jal _ _ => Return Jal_leakage | InvalidI => Return InvalidI_leakage @@ -422,8 +425,8 @@ Inductive InstructionF : Type := Definition leakage_of_instr_F (instr : InstructionF) : M LeakageF := match instr with - | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (rs1_val + oimm12)) - | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (rs1_val + simm12)) + | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (regToZ_unsigned rs1_val + oimm12)) + | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (regToZ_unsigned rs1_val + simm12)) | Fmadd_s _ _ _ _ _ => Return Fmadd_s_leakage | Fmsub_s _ _ _ _ _ => Return Fmsub_s_leakage | Fnmsub_s _ _ _ _ _ => Return Fnmsub_s_leakage @@ -564,17 +567,17 @@ Inductive InstructionA64 : Type := Definition leakage_of_instr_A64 (instr : InstructionA64) : M LeakageA64 := match instr with - | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage rs1_val) - | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage rs1_val) - | Amoswap_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_d_leakage rs1_val) - | Amoadd_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_d_leakage rs1_val) - | Amoand_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_d_leakage rs1_val) - | Amoor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_d_leakage rs1_val) - | Amoxor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_d_leakage rs1_val) - | Amomax_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_d_leakage rs1_val) - | Amomaxu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_d_leakage rs1_val) - | Amomin_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_d_leakage rs1_val) - | Amominu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_d_leakage rs1_val) + | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage (regToZ_unsigned rs1_val)) + | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage (regToZ_unsigned rs1_val)) + | Amoswap_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_d_leakage (regToZ_unsigned rs1_val)) + | Amoadd_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_d_leakage (regToZ_unsigned rs1_val)) + | Amoand_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_d_leakage (regToZ_unsigned rs1_val)) + | Amoor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_d_leakage (regToZ_unsigned rs1_val)) + | Amoxor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_d_leakage (regToZ_unsigned rs1_val)) + | Amomax_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_d_leakage (regToZ_unsigned rs1_val)) + | Amomaxu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_d_leakage (regToZ_unsigned rs1_val)) + | Amomin_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_d_leakage (regToZ_unsigned rs1_val)) + | Amominu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_d_leakage (regToZ_unsigned rs1_val)) | InvalidA64 => Return InvalidA64_leakage end. @@ -629,17 +632,17 @@ Inductive InstructionA : Type := Definition leakage_of_instr_A (instr : InstructionA) : M LeakageA := match instr with - | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage rs1_val) - | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage rs1_val) - | Amoswap_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_w_leakage rs1_val) - | Amoadd_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_w_leakage rs1_val) - | Amoand_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_w_leakage rs1_val) - | Amoor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_w_leakage rs1_val) - | Amoxor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_w_leakage rs1_val) - | Amomax_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_w_leakage rs1_val) - | Amomaxu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_w_leakage rs1_val) - | Amomin_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_w_leakage rs1_val) - | Amominu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_w_leakage rs1_val) + | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage (regToZ_unsigned rs1_val)) + | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage (regToZ_unsigned rs1_val)) + | Amoswap_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_w_leakage (regToZ_unsigned rs1_val)) + | Amoadd_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_w_leakage (regToZ_unsigned rs1_val)) + | Amoand_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_w_leakage (regToZ_unsigned rs1_val)) + | Amoor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_w_leakage (regToZ_unsigned rs1_val)) + | Amoxor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_w_leakage (regToZ_unsigned rs1_val)) + | Amomax_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_w_leakage (regToZ_unsigned rs1_val)) + | Amomaxu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_w_leakage (regToZ_unsigned rs1_val)) + | Amomin_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_w_leakage (regToZ_unsigned rs1_val)) + | Amominu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_w_leakage (regToZ_unsigned rs1_val)) | InvalidA => Return InvalidA_leakage end. From f4d77bd6861c9274d9009501b629df185ce571ba Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 5 Nov 2023 16:23:55 -0500 Subject: [PATCH 08/42] updated run1 --- src/riscv/Platform/Run.v | 14 +++++++++----- 1 file changed, 9 insertions(+), 5 deletions(-) diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index ec56241..ea86151 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -15,16 +15,20 @@ Section Riscv. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVP: RiscvProgramWithLeakage}. + Context {RVM: RiscvProgramWithLeakage}. Context {RVS: RiscvMachine M mword}. + Print leakage_of_instr. Print decode. Print leakEvent. Print getRegister. Print MachineWidth. + Definition run1(iset: InstructionSet): M unit := pc <- getPC; - inst <- loadWord Fetch pc; - logInstr (decode iset (combine 4 inst));; - execute (decode iset (combine 4 inst));; - endCycleNormal. + inst <- loadWord Fetch pc; + let inst' := decode iset (combine 4 inst) in + leakage_event <- leakage_of_instr getRegister inst'; + leakEvent leakage_event;; + execute inst';; + endCycleNormal. (* Note: We cannot use power_func (fun m => run1 iset;; m) n (Return tt) From dbc592d330391f2354deab31a3616a9f8e4e88a9 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 5 Nov 2023 16:29:58 -0500 Subject: [PATCH 09/42] made MinimalMMIO build with leakEvent thing --- src/riscv/Platform/MaterializeRiscvProgram.v | 4 ++-- src/riscv/Platform/MetricMinimalMMIO.v | 2 +- src/riscv/Platform/MinimalMMIO.v | 19 ++----------------- 3 files changed, 5 insertions(+), 20 deletions(-) diff --git a/src/riscv/Platform/MaterializeRiscvProgram.v b/src/riscv/Platform/MaterializeRiscvProgram.v index 386aa76..f49280e 100644 --- a/src/riscv/Platform/MaterializeRiscvProgram.v +++ b/src/riscv/Platform/MaterializeRiscvProgram.v @@ -27,7 +27,7 @@ Section Riscv. | GetPrivMode | SetPrivMode (_ : PrivMode) | Fence (_ : MachineInt) (_ : MachineInt) - | LogInstr (_ : Instruction) + | LeakEvent (_ : LeakageEvent) | GetPC | SetPC (_ : word) | StartCycle @@ -53,7 +53,7 @@ Section Riscv. | GetPrivMode => PrivMode | SetPrivMode _ => unit | Fence _ _ => unit - | LogInstr _ => unit + | LeakEvent _ => unit | GetPC => word | SetPC _ => unit | StartCycle => unit diff --git a/src/riscv/Platform/MetricMinimalMMIO.v b/src/riscv/Platform/MetricMinimalMMIO.v index 8d31115..1090c2a 100644 --- a/src/riscv/Platform/MetricMinimalMMIO.v +++ b/src/riscv/Platform/MetricMinimalMMIO.v @@ -64,7 +64,7 @@ Section Riscv. Global Instance IsRiscvMachineWithLeakage: RiscvProgramWithLeakage := {| - logInstr a := act (id, LogInstr a) ret; |}. + leakEvent a := act (id, LeakEvent a) ret; |}. Definition interp_action a metmach post := interpret_action (snd a) (metmach.(getMachine)) (fun r mach => diff --git a/src/riscv/Platform/MinimalMMIO.v b/src/riscv/Platform/MinimalMMIO.v index 3f74776..8f046df 100644 --- a/src/riscv/Platform/MinimalMMIO.v +++ b/src/riscv/Platform/MinimalMMIO.v @@ -22,20 +22,6 @@ Require Import riscv.Platform.Sane. Local Open Scope Z_scope. Local Open Scope bool_scope. -Definition Mtriv (x : Type) := x. -Definition trivialBind (A B : Type) (x : Mtriv A) (f : A -> B) : Mtriv B := f x. -Definition trivialReturn (A : Type) (a : A) : Mtriv A := a. -Print Monad. -Lemma trivial_left_identity : forall (A B : Type) (a : A) (f : A -> Mtriv B), trivialBind A B (trivialReturn A a) f = f a. -Proof. trivial. Qed. -Lemma trivial_right_identity : forall (A : Type) (m : Mtriv A), trivialBind A A m (trivialReturn A) = m. -Proof. trivial. Qed. -Lemma trivial_associativity : forall (A B C : Type) (m : Mtriv A) (f : A -> Mtriv B) (g : B -> Mtriv C), trivialBind B C (trivialBind A B m f) g = trivialBind A C m (fun x : A => trivialBind B C (f x) g). -Proof. trivial. Qed. - -Definition trivialMonad : Monad Mtriv := - {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. - Class MMIOSpec{width: Z}{BW: Bitwidth width}{word: word width}{Mem : map.map word byte} := { (* should not say anything about alignment, just whether it's in the MMIO range *) isMMIOAddr: word -> Prop; @@ -103,7 +89,7 @@ Section Riscv. else word.of_Z 0. Definition setReg(reg: Z)(v: word)(regs: Registers): Registers := - if ((0 RiscvMachine -> Prop) -> (RiscvMachine -> Prop) -> Prop := @@ -128,8 +114,7 @@ Section Riscv. postF tt (withNextPc (word.add mach.(getPc) (word.of_Z 4)) mach) | EndCycleNormal => fun postF postA => postF tt (updatePc mach) | EndCycleEarly _ => fun postF postA => postA (updatePc mach) (* ignores postF containing the continuation *) - | LogInstr i => fun (postF : unit -> RiscvMachine -> Prop) postA => - postF tt (withTraceItem (@leakage_of_instr Mtriv trivialMonad (fun reg => word.unsigned (getReg mach.(getRegs) reg)) i) mach) + | LeakEvent e => fun postF postA => postF tt (withLeakageEvent e mach) | MakeReservation _ | ClearReservation _ | CheckReservation _ From 6d372e5bdaf7c6db8670fe3c226ccf9676b8ecaf Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 5 Nov 2023 20:45:20 -0500 Subject: [PATCH 10/42] started trying to make the thing build. giving the default "False spec" for LeakEvent in MinimalNoMul was no good; proof in MetricMinimalNoMul doesn't go through. Will need to fix that. --- src/riscv/Platform/MaterializeRiscvProgram.v | 7 +- src/riscv/Platform/MetricMinimalNoMul.v | 147 ++++++++ src/riscv/Platform/MetricSane.v | 11 +- src/riscv/Platform/Minimal.v | 21 +- src/riscv/Platform/MinimalNoMul.v | 370 +++++++++++++++++++ src/riscv/Spec/MetricPrimitives.v | 2 +- src/riscv/Spec/Primitives.v | 21 +- 7 files changed, 564 insertions(+), 15 deletions(-) create mode 100644 src/riscv/Platform/MetricMinimalNoMul.v create mode 100644 src/riscv/Platform/MinimalNoMul.v diff --git a/src/riscv/Platform/MaterializeRiscvProgram.v b/src/riscv/Platform/MaterializeRiscvProgram.v index f49280e..36de0e1 100644 --- a/src/riscv/Platform/MaterializeRiscvProgram.v +++ b/src/riscv/Platform/MaterializeRiscvProgram.v @@ -84,7 +84,12 @@ Section Riscv. setPC a := act (SetPC a) ret; endCycleNormal := act EndCycleNormal ret; endCycleEarly A := act (EndCycleEarly A) ret; - |}. + |}. + + Global Instance MaterializeWithLeakage : RiscvProgramWithLeakage := + {| + RVP := Materialize; + leakEvent a := act (LeakEvent a) ret |}. (* Not (yet) in Riscv monad, but added here because it's useful to initialize nextPc to pc+4 at the beginning of each cycle instead of having to maintain diff --git a/src/riscv/Platform/MetricMinimalNoMul.v b/src/riscv/Platform/MetricMinimalNoMul.v new file mode 100644 index 0000000..1c2acfb --- /dev/null +++ b/src/riscv/Platform/MetricMinimalNoMul.v @@ -0,0 +1,147 @@ +Require Import Coq.ZArith.ZArith. +Require Import Coq.Logic.FunctionalExtensionality. +Require Import Coq.Logic.PropExtensionality. +Require Import riscv.Utility.Monads. +Require Import riscv.Utility.MonadNotations. +Require Import riscv.Spec.Decode. +Require Import riscv.Spec.Machine. +Require Import riscv.Utility.Utility. +Require Import riscv.Spec.Primitives. +Require Import riscv.Spec.MetricPrimitives. +Require Import Coq.Lists.List. Import ListNotations. +Require Export riscv.Platform.RiscvMachine. +Require Export riscv.Platform.MetricRiscvMachine. +Require Import riscv.Platform.MinimalNoMul. +Require Import riscv.Platform.MetricLogging. +Require Import coqutil.Z.Lia. +Require Import coqutil.Map.Interface. +Require Import coqutil.Tactics.Tactics. +Import MetricRiscvMachine. + +Local Open Scope Z_scope. +Local Open Scope bool_scope. + +Section Riscv. + Import List. + Import free. + + Context {width: Z} {BW: Bitwidth width} {word: word width} {word_ok: word.ok word}. + Context {Mem: map.map word byte}. + Context {Registers: map.map Register word}. + + Local Notation M := (free riscv_primitive primitive_result). + + Definition metrics_of(p: riscv_primitive): MetricLog -> MetricLog := + match p with + | GetRegister a => id + | SetRegister a b => id + | LoadByte a b => addMetricLoads 1 + | LoadHalf a b => addMetricLoads 1 + | LoadWord a b => addMetricLoads 1 + | LoadDouble a b => addMetricLoads 1 + | StoreByte a b c => addMetricStores 1 + | StoreHalf a b c => addMetricStores 1 + | StoreWord a b c => addMetricStores 1 + | StoreDouble a b c => addMetricStores 1 + | MakeReservation a => id + | ClearReservation a => id + | CheckReservation a => id + | GetCSRField f => id + | SetCSRField f v => id + | GetPrivMode => id + | SetPrivMode m => id + | Fence a b => id + | LeakEvent a => id + | GetPC => id + | SetPC a => addMetricJumps 1 + | StartCycle => id + | EndCycleNormal => addMetricInstructions 1 + | EndCycleEarly A => addMetricInstructions 1 + end. + + Definition interp_action p (m: MetricRiscvMachine) post := + interpret_action p m.(getMachine) + (fun r mach => post r (mkMetricRiscvMachine mach (metrics_of p m.(getMetrics)))) + (fun _ => False). + + Arguments Memory.load_bytes: simpl never. + Arguments Memory.store_bytes: simpl never. + Arguments LittleEndian.combine: simpl never. + + Global Instance MetricMinimalNoMulPrimitivesParams: PrimitivesParams M MetricRiscvMachine := + { + Primitives.mcomp_sat := @free.interp _ _ _ interp_action; + Primitives.is_initial_register_value x := True; + Primitives.nonmem_load := Primitives.nonmem_load (PrimitivesParams := MinimalNoMulPrimitivesParams); + Primitives.nonmem_store := Primitives.nonmem_store (PrimitivesParams := MinimalNoMulPrimitivesParams); + Primitives.valid_machine mach := no_M mach.(getMachine); + }. + + Global Instance MinimalNoMulSatisfies_mcomp_sat_spec: mcomp_sat_spec MetricMinimalNoMulPrimitivesParams. + Proof. + split; cbv [mcomp_sat MetricMinimalNoMulPrimitivesParams Monad_free Bind Return]. + { symmetry. eapply interp_bind_ex_mid; intros. + eapply MinimalNoMul.interpret_action_weaken_post; eauto; cbn; eauto. } + { symmetry. rewrite interp_ret; eapply iff_refl. } + Qed. + + Lemma interp_action_weaken_post a (post1 post2:_->_->Prop) + (H: forall r s, post1 r s -> post2 r s) s + : interp_action a s post1 -> interp_action a s post2. + Proof. eapply MinimalNoMul.interpret_action_weaken_post; eauto. Qed. + Lemma interp_action_appendonly' a s post : + interp_action a s post -> + interp_action a s (fun v s' => post v s' /\ endswith s'.(getLog) s.(getLog)). + Proof. eapply MinimalNoMul.interpret_action_appendonly''; eauto. Qed. + Lemma interp_action_total{memOk: map.ok Mem} a (s: MetricRiscvMachine) post : + no_M s -> + interp_action a s post -> + exists v s, post v s /\ no_M s. + Proof. + intros H H1. + unshelve epose proof (MinimalNoMul.interpret_action_total _ _ _ _ _ H1) as H0; eauto. + destruct H0 as (?&?&[[]|(?&?)]); eauto. + Qed. + Lemma interp_action_preserves_valid{memOk: map.ok Mem} a s post : + no_M s.(getMachine) -> + interp_action a s post -> + interp_action a s (fun v s' => post v s' /\ no_M s'.(getMachine)). + Proof. + intros D I. + unshelve epose proof (MinimalNoMul.interpret_action_preserves_valid' _ _ _ D I) as H0; eauto. + Qed. Search RiscvProgramWithLeakage. Print MetricPrimitivesSane. Search RiscvProgramWithLeakage. + + Global Instance MetricMinimalNoMulPrimitivesSane{memOk: map.ok Mem} : + MetricPrimitivesSane MetricMinimalNoMulPrimitivesParams. + Proof. + split; cbv [mcomp_sane valid_machine MetricMinimalNoMulPrimitivesParams]; + intros *; intros D M; + (split; [ exact (interp_action_total _ st _ D M) + | eapply interp_action_preserves_valid; try eassumption; + eapply interp_action_appendonly'; try eassumption ]). + Qed. + + Global Instance MetricMinimalNoMulSatisfiesPrimitives{memOk: map.ok Mem}: + MetricPrimitives MetricMinimalNoMulPrimitivesParams. + Proof. + split; try exact _. + all : cbv [mcomp_sat spec_load spec_store MetricMinimalNoMulPrimitivesParams invalidateWrittenXAddrs]. + all: intros; destruct initialL; + repeat match goal with + | _ => progress subst + | _ => Option.inversion_option + | _ => progress cbn -[Memory.load_bytes Memory.store_bytes HList.tuple] in * + | _ => progress cbv [id valid_register is_initial_register_value load store Memory.loadByte Memory.loadHalf Memory.loadWord Memory.loadDouble Memory.storeByte Memory.storeHalf Memory.storeWord Memory.storeDouble] in * + | H : exists _, _ |- _ => destruct H + | H : _ /\ _ |- _ => destruct H + | |- _ => solve [ intuition (eauto || blia) ] + | H : _ \/ _ |- _ => destruct H + | |- context[match ?x with _ => _ end] => destruct x eqn:? + | |- _ => progress unfold getReg, setReg + | |-_ /\ _ => split + end. + (* setRegister *) + 1: destruct getMachine; eassumption. + Qed. + +End Riscv. diff --git a/src/riscv/Platform/MetricSane.v b/src/riscv/Platform/MetricSane.v index 05d75e1..5e5fde1 100644 --- a/src/riscv/Platform/MetricSane.v +++ b/src/riscv/Platform/MetricSane.v @@ -151,6 +151,14 @@ Section Sane. Ltac t := repeat (simpl; unfold when; repeat t_step'). + Lemma leakage_of_instr_sane: forall inst, + mcomp_sane (leakage_of_instr getRegister inst). + Proof. + intros. + destruct inst as [inst | inst | inst | inst | inst | inst | inst | inst | inst | inst]; + simpl; destruct inst; t. + Qed. + Lemma execute_sane: forall inst, mcomp_sane (Execute.execute inst). Proof. @@ -165,7 +173,8 @@ Section Sane. unfold run1. intros. apply Bind_sane; [apply getPC_sane|intros]. apply Bind_sane; [apply loadWord_sane|intros]. - apply Bind_sane; [apply logInstr_sane|intros]. + apply Bind_sane; [apply leakage_of_instr_sane|intros]. + apply Bind_sane; [apply leakInstr_sane|intros]. apply Bind_sane; [apply execute_sane|intros]. apply endCycleNormal_sane. Qed. diff --git a/src/riscv/Platform/Minimal.v b/src/riscv/Platform/Minimal.v index a2da920..5f49695 100644 --- a/src/riscv/Platform/Minimal.v +++ b/src/riscv/Platform/Minimal.v @@ -98,7 +98,13 @@ Section Riscv. (* fail hard if exception is thrown because at the moment, we want to prove that code output by the compiler never throws exceptions *) endCycleEarly{A: Type} := fail_hard; - }. + }. + + Print RiscvProgramWithLeakage. + Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage (OState RiscvMachine) _ _ _ := { + RVP := IsRiscvMachine; + leakEvent e := fail_hard; + }. Arguments Memory.load_bytes: simpl never. Arguments Memory.store_bytes: simpl never. @@ -116,7 +122,7 @@ Section Riscv. | |- _ => reflexivity | |- _ => progress ( unfold computation_satisfies, computation_with_answer_satisfies, - IsRiscvMachine, + IsRiscvMachine, IsRiscvMachineWithLeakage, valid_register, is_initial_register_value, get, put, fail_hard, @@ -225,10 +231,20 @@ Section Riscv. intros. eapply update_sane. intros. exists [e]. destruct mach. reflexivity. Qed. + Print PrimitivesSane. + Instance MinimalSane: PrimitivesSane MinimalPrimitivesParams. Proof. constructor. all: intros; + unfold IsRiscvMachine, IsRiscvMachineWithLeakage, RVP, + getRegister, setRegister, + loadByte, loadHalf, loadWord, loadDouble, + storeByte, storeHalf, storeWord, storeDouble, + getPC, setPC, + endCycleNormal, endCycleEarly, raiseExceptionWithInfo, + loadN, storeN, fail_if_None. + all: unfold getRegister, setRegister, loadByte, loadHalf, loadWord, loadDouble, storeByte, storeHalf, storeWord, storeDouble, @@ -261,4 +277,5 @@ End Riscv. (* needed because defined inside a Section *) #[global] Existing Instance IsRiscvMachine. +#[global] Existing Instance IsRiscvMachineWithLeakage. #[global] Existing Instance MinimalSatisfiesPrimitives. diff --git a/src/riscv/Platform/MinimalNoMul.v b/src/riscv/Platform/MinimalNoMul.v new file mode 100644 index 0000000..d357e33 --- /dev/null +++ b/src/riscv/Platform/MinimalNoMul.v @@ -0,0 +1,370 @@ +Require Import Coq.Strings.String. +Require Import Coq.ZArith.ZArith. +Require Import riscv.Utility.Monads. +Require Import riscv.Utility.MonadNotations. +Require Export riscv.Utility.FreeMonad. +Require Import riscv.Spec.Decode. +Require Import riscv.Spec.Machine. +Require Import riscv.Utility.Utility. +Require Import riscv.Spec.Primitives. +Require Import Coq.Lists.List. Import ListNotations. +Require Import coqutil.Datatypes.List. +Require Import coqutil.Datatypes.ListSet. +Require Export riscv.Platform.RiscvMachine. +Require Export riscv.Platform.MaterializeRiscvProgram. +Require Import coqutil.Z.Lia. +Require Import coqutil.Map.Interface. +Require Import coqutil.Map.Properties. +Require Import coqutil.Word.Properties. +Require Import coqutil.Datatypes.PropSet. +Require Import coqutil.Tactics.Tactics. +Require Import coqutil.Tactics.fwd. +Require Import riscv.Platform.Sane. + +Local Open Scope Z_scope. +Local Open Scope bool_scope. + +Section Riscv. + Import free. + Context {width: Z} {BW: Bitwidth width} {word: word width} {word_ok: word.ok word}. + Context {Mem: map.map word byte} {Registers: map.map Register word}. + + Add Ring wring : (word.ring_theory (word := word)) + (preprocess [autorewrite with rew_word_morphism], + morphism (word.ring_morph (word := word)), + constants [word_cst]). + + Definition store(n: nat)(ctxid: SourceType) a v mach post := + match Memory.store_bytes n mach.(getMem) a v with + | Some m => post (withXAddrs (invalidateWrittenXAddrs n a mach.(getXAddrs)) (withMem m mach)) + | None => False + end. + + Definition load(n: nat)(ctxid: SourceType) a mach post := + (ctxid = Fetch -> isXAddr4 a mach.(getXAddrs)) /\ + match Memory.load_bytes n mach.(getMem) a with + | Some v => post v mach + | None => False + end. + + Definition updatePc(mach: RiscvMachine): RiscvMachine := + withPc mach.(getNextPc) (withNextPc (word.add mach.(getNextPc) (word.of_Z 4)) mach). + + Definition getReg(regs: Registers)(reg: Z): word := + if ((0 x + | None => word.of_Z 0 + end + else word.of_Z 0. + + Definition setReg(reg: Z)(v: word)(regs: Registers): Registers := + if ((0 RiscvMachine -> Prop) -> (RiscvMachine -> Prop) -> Prop := + match a with + | GetRegister reg => fun postF postA => + let v := getReg mach.(getRegs) reg in + postF v mach + | SetRegister reg v => fun postF postA => + let regs := setReg reg v mach.(getRegs) in + postF tt (withRegs regs mach) + | GetPC => fun postF postA => postF mach.(getPc) mach + | SetPC newPC => fun postF postA => postF tt (withNextPc newPC mach) + | LoadByte ctxid a => fun postF postA => load 1 ctxid a mach postF + | LoadHalf ctxid a => fun postF postA => load 2 ctxid a mach postF + | LoadWord ctxid a => fun postF postA => load 4 ctxid a mach postF + | LoadDouble ctxid a => fun postF postA => load 8 ctxid a mach postF + | StoreByte ctxid a v => fun postF postA => store 1 ctxid a v mach (postF tt) + | StoreHalf ctxid a v => fun postF postA => store 2 ctxid a v mach (postF tt) + | StoreWord ctxid a v => fun postF postA => store 4 ctxid a v mach (postF tt) + | StoreDouble ctxid a v => fun postF postA => store 8 ctxid a v mach (postF tt) + | StartCycle => fun postF postA => + postF tt (withNextPc (word.add mach.(getPc) (word.of_Z 4)) mach) + | EndCycleNormal => fun postF postA => postF tt (updatePc mach) + | EndCycleEarly _ => fun postF postA => postA (updatePc mach) (* ignores postF containing the continuation *) + | LeakEvent _ + | MakeReservation _ + | ClearReservation _ + | CheckReservation _ + | GetCSRField _ + | SetCSRField _ _ + | GetPrivMode + | SetPrivMode _ + | Fence _ _ + => fun postF postA => False + end. + + Definition no_M(mach: RiscvMachine): Prop := + forall a v, + isXAddr4 a mach.(getXAddrs) -> + word.unsigned a mod 4 = 0 -> + Memory.load_bytes 4 mach.(getMem) a = Some v -> + forall minst, decode RV32IM (LittleEndian.combine 4 v) <> MInstruction minst. + + Instance MinimalNoMulPrimitivesParams: + PrimitivesParams (free riscv_primitive primitive_result) RiscvMachine := + {| + Primitives.mcomp_sat A m mach postF := + @free.interpret _ _ _ interpret_action A m mach postF (fun _ => False); + Primitives.is_initial_register_value x := True; + Primitives.nonmem_load _ _ _ _ _ := False; + Primitives.nonmem_store _ _ _ _ _ _ := False; + Primitives.valid_machine := no_M; + |}. + + Lemma load_weaken_post n c a m (post1 post2:_->_->Prop) + (H: forall r s, post1 r s -> post2 r s) + : load n c a m post1 -> load n c a m post2. + Proof. + cbv [load nonmem_load]. + destruct (Memory.load_bytes n (getMem m) a); intuition eauto. + Qed. + + Lemma store_weaken_post n c a v m (post1 post2:_->Prop) + (H: forall s, post1 s -> post2 s) + : store n c a v m post1 -> store n c a v m post2. + Proof. + cbv [store nonmem_store]. + destruct (Memory.store_bytes n (getMem m) a); intuition eauto. + Qed. + + Lemma interpret_action_weaken_post a (postF1 postF2: _ -> _ -> Prop) (postA1 postA2: _ -> Prop): + (forall r s, postF1 r s -> postF2 r s) -> + (forall s, postA1 s -> postA2 s) -> + forall s, interpret_action a s postF1 postA1 -> interpret_action a s postF2 postA2. + Proof. + destruct a; cbn; try solve [intuition eauto]. + all : eauto using load_weaken_post, store_weaken_post. + Qed. + + Global Instance MinimalNoMulSatisfies_mcomp_sat_spec: mcomp_sat_spec MinimalNoMulPrimitivesParams. + Proof. + split; cbv [mcomp_sat MinimalNoMulPrimitivesParams Bind Return Monad_free]. + { symmetry. eapply interpret_bind_ex_mid, interpret_action_weaken_post. } + { symmetry; intros. rewrite interpret_ret; eapply iff_refl. } + Qed. + + Lemma preserve_undef_on{memOk: map.ok Mem}: forall n (m m': Mem) a w s, + Memory.store_bytes n m a w = Some m' -> + map.undef_on m s -> + map.undef_on m' s. + Proof. + eauto using map.same_domain_preserves_undef_on, Memory.store_bytes_preserves_domain. + Qed. + + Lemma removeXAddr_bw: forall (a1 a2: word) xaddrs, + isXAddr1 a1 (removeXAddr a2 xaddrs) -> + isXAddr1 a1 xaddrs. + Proof. + unfold isXAddr1, removeXAddr. + intros. + eapply filter_In. + eassumption. + Qed. + + Lemma invalidateWrittenXAddrs_bw: forall n (a r: word) xa, + In a (invalidateWrittenXAddrs n r xa) -> + In a xa. + Proof. + induction n; cbn; intros. + - assumption. + - eapply IHn. eapply removeXAddr_bw. unfold isXAddr1. eassumption. + Qed. + + Lemma put_preserves_getmany_of_tuple{memOk: map.ok Mem}: + forall n (t: HList.tuple word n) (m: Mem) (r: word) b, + ~In r (HList.tuple.to_list t) -> + map.getmany_of_tuple m t = + map.getmany_of_tuple (map.put m r b) t. + Proof. + induction n; intros. + - destruct t. reflexivity. + - destruct t as (w & t). cbn in H|-*. + unfold map.getmany_of_tuple in IHn. + erewrite IHn. + 2: { + intro C. eapply H. right. exact C. + } + rewrite ?map.get_put_dec. + destr (word.eqb r w). 2: reflexivity. + exfalso. apply H. auto. + Qed. + + Lemma transfer_load4bytes_to_previous_mem{memOk: map.ok Mem}: + forall n (a: word) v m m' r w someSet, + Memory.store_bytes n m r w = Some m' -> + (* a notin r..r+n *) + isXAddr4 a (invalidateWrittenXAddrs n r someSet) -> + Memory.load_bytes 4 m' a = Some v -> + Memory.load_bytes 4 m a = Some v. + Proof. + induction n; intros. + - cbn in H. congruence. + - unfold Memory.store_bytes in *. fwd. cbn in H0. destruct w as [b w]. + cbn -[HList.tuple Memory.load_bytes] in H1. + cbn in E. fwd. + unfold Memory.load_bytes at 1 in IHn. + unfold map.getmany_of_tuple, Memory.footprint in IHn. + specialize IHn with (m := m) (r := (word.add r (word.of_Z 1))). + rewrite E1 in IHn. + eapply IHn. + + reflexivity. + + instantiate (1 := someSet). clear -H0. + unfold isXAddr4 in *. fwd. eauto 10 using removeXAddr_bw. + + unfold Memory.load_bytes in *. + etransitivity. 2: eassumption. + unfold Memory.unchecked_store_bytes, Memory.footprint. + eapply put_preserves_getmany_of_tuple. + cbn. clear -H0 word_ok. unfold isXAddr4, isXAddr1 in H0. fwd. + intro C. + unfold removeXAddr in *. + apply_in_hyps filter_In. + fwd. + apply_in_hyps Bool.negb_true_iff. + apply_in_hyps Properties.word.eqb_false. + repeat destruct C as [C | C]; try assumption; + match type of C with + | ?l = _ => ring_simplify l in C + end; + subst r; + congruence. + Qed. + + Lemma isXAddr4_uninvalidate: forall (a: word) n r xa, + isXAddr4 a (invalidateWrittenXAddrs n r xa) -> + isXAddr4 a xa. + Proof. + unfold isXAddr4, isXAddr1. intros. fwd. eauto 10 using invalidateWrittenXAddrs_bw. + Qed. + + Lemma interpret_action_total{memOk: map.ok Mem} a s postF postA : + no_M s -> + interpret_action a s postF postA -> + exists s', no_M s' /\ (postA s' \/ exists v', postF v' s'). + Proof. + destruct s, a; cbn -[HList.tuple Memory.load_bytes invalidateWrittenXAddrs]; + cbv [load store no_M]; cbn -[HList.tuple Memory.load_bytes invalidateWrittenXAddrs]; + repeat destruct_one_match; + intuition idtac; + repeat lazymatch goal with + | H : postF _ ?mach |- exists _ : RiscvMachine, _ => + exists mach; cbn [RiscvMachine.getMem RiscvMachine.getXAddrs] + | H : postA ?mach |- exists _ : RiscvMachine, _ => + exists mach; cbn [RiscvMachine.getMem RiscvMachine.getXAddrs] + | Hexists : (exists v, ?P), Hforall : (forall v, ?P -> _) |- _ => + let v := fresh "v" in + destruct Hexists as [v Hexists]; + specialize (Hforall v Hexists) + end; + ssplit; eauto; cbn -[HList.tuple Memory.load_bytes invalidateWrittenXAddrs]; + change removeXAddr with (@List.removeb word word.eqb); + rewrite ?ListSet.of_list_removeb; + intuition eauto 10 using transfer_load4bytes_to_previous_mem, isXAddr4_uninvalidate. + Qed. + + Lemma interpret_action_total'{memOk: map.ok Mem} a s post : + no_M s -> + interpret_action a s post (fun _ : RiscvMachine => False) -> + exists v s', post v s' /\ no_M s'. + Proof. + intros. pose proof interpret_action_total as P. + specialize P with (postA := (fun _ : RiscvMachine => False)). simpl in P. + specialize (P _ _ _ H H0). + destruct P as (s' & ? & ?). + destruct H2 as [[] | (v' & ?)]. + eauto. + Qed. + + Import coqutil.Tactics.Tactics. + + Lemma interpret_action_appendonly a s postF postA : + interpret_action a s postF postA -> + interpret_action a s (fun _ s' => endswith s'.(getLog) s.(getLog)) + (fun s' => endswith s'.(getLog) s.(getLog)). + Proof. + destruct s, a; cbn; cbv [load store nonmem_load nonmem_store]; cbn; + repeat destruct_one_match; + intuition eauto using endswith_refl, endswith_cons_l. + Qed. + + (* NOTE: maybe instead a generic lemma to push /\ into postcondition? *) + Lemma interpret_action_appendonly' a s postF postA : + interpret_action a s postF postA -> + interpret_action a s (fun v s' => postF v s' /\ endswith s'.(getLog) s.(getLog)) + (fun s' => postA s' /\ endswith s'.(getLog) s.(getLog)). + Proof. + destruct s, a; cbn; cbv [load store nonmem_load nonmem_store]; cbn; + repeat destruct_one_match; intros; destruct_products; try split; + intuition eauto using endswith_refl, endswith_cons_l. + Qed. + + Lemma interpret_action_appendonly'' a s post : + interpret_action a s post (fun _ : RiscvMachine => False) -> + interpret_action a s (fun v s' => post v s' /\ endswith s'.(getLog) s.(getLog)) + (fun _ : RiscvMachine => False). + Proof. + intros. pose proof interpret_action_appendonly' as P. + specialize (P _ _ _ (fun _ : RiscvMachine => False) H). simpl in P. + eapply interpret_action_weaken_post. 3: exact P. all: simpl; intuition eauto. + Qed. + + Lemma interpret_action_preserves_valid{memOk: map.ok Mem} a s postF postA : + no_M s -> + interpret_action a s postF postA -> + interpret_action a s (fun v s' => postF v s' /\ no_M s') + (fun s' => postA s' /\ no_M s'). + Proof. + destruct s, a; cbn; cbv [load store no_M]; + cbn -[HList.tuple Memory.load_bytes invalidateWrittenXAddrs]; + repeat destruct_one_match; intros; destruct_products; try split; + change removeXAddr with (@List.removeb word word.eqb); + rewrite ?ListSet.of_list_removeb; + intuition eauto 10 using transfer_load4bytes_to_previous_mem, isXAddr4_uninvalidate. + Qed. + + Lemma interpret_action_preserves_valid'{memOk: map.ok Mem} a s post : + no_M s -> + interpret_action a s post (fun _ : RiscvMachine => False) -> + interpret_action a s (fun v s' => post v s' /\ no_M s') + (fun _ : RiscvMachine => False). + Proof. + intros. pose proof interpret_action_preserves_valid as P. + specialize (P _ _ _ (fun _ : RiscvMachine => False) H H0). simpl in P. + eapply interpret_action_weaken_post. 3: exact P. all: simpl; intuition eauto. + Qed. + + Global Instance MinimalNoMulPrimitivesSane{memOk: map.ok Mem} : + PrimitivesSane MinimalNoMulPrimitivesParams. + Proof. + split; cbv [mcomp_sane valid_machine MinimalNoMulPrimitivesParams]; intros *; intros D M; + (split; [ exact (interpret_action_total' _ st _ D M) + | eapply interpret_action_preserves_valid'; try eassumption; + eapply interpret_action_appendonly''; try eassumption ]). + Qed. + + Global Instance MinimalNoMulSatisfiesPrimitives{memOk: map.ok Mem} : + Primitives MinimalNoMulPrimitivesParams. + Proof. + split; try exact _. + all : cbv [mcomp_sat spec_load spec_store MinimalNoMulPrimitivesParams invalidateWrittenXAddrs]. + all: intros; + repeat match goal with + | _ => progress subst + | _ => Option.inversion_option + | _ => progress cbn -[Memory.load_bytes Memory.store_bytes HList.tuple] + | _ => progress cbv [valid_register is_initial_register_value load store Memory.loadByte Memory.loadHalf Memory.loadWord Memory.loadDouble Memory.storeByte Memory.storeHalf Memory.storeWord Memory.storeDouble] in * + | H : exists _, _ |- _ => destruct H + | H : _ /\ _ |- _ => destruct H + | |- _ => solve [ intuition (eauto || blia) ] + | H : _ \/ _ |- _ => destruct H + | |- context[match ?x with _ => _ end] => destruct x eqn:? + | |- _ => progress unfold getReg, setReg + | |-_ /\ _ => split + end. + (* setRegister *) + destruct initialL; eassumption. + Qed. + +End Riscv. diff --git a/src/riscv/Spec/MetricPrimitives.v b/src/riscv/Spec/MetricPrimitives.v index 29c212b..3e913a4 100644 --- a/src/riscv/Spec/MetricPrimitives.v +++ b/src/riscv/Spec/MetricPrimitives.v @@ -55,7 +55,7 @@ Section MetricPrimitives. getPrivMode_sane: mcomp_sane getPrivMode; setPrivMode_sane: forall m, mcomp_sane (setPrivMode m); fence_sane: forall a b, mcomp_sane (fence a b); - leakInstr_sane: forall a, mcomp_sane (leakEvent a); + leakEvent_sane: forall a, mcomp_sane (leakEvent a); getPC_sane: mcomp_sane getPC; setPC_sane: forall newPc, mcomp_sane (setPC newPc); endCycleNormal_sane: mcomp_sane endCycleNormal; diff --git a/src/riscv/Spec/Primitives.v b/src/riscv/Spec/Primitives.v index a04067d..09b5a5c 100644 --- a/src/riscv/Spec/Primitives.v +++ b/src/riscv/Spec/Primitives.v @@ -70,8 +70,8 @@ Section Primitives. (mcomp_sat comp st (fun a st' => (post a st' /\ exists diff, st'.(getLog) = diff ++ st.(getLog)) /\ valid_machine st')). - Context {RVM: RiscvProgram M word}. - Context {RVS: @riscv.Spec.Machine.RiscvMachine M word _ _ RVM}. + Context {RVM: RiscvProgramWithLeakage}. + Context {RVS: @riscv.Spec.Machine.RiscvMachine M word _ _ RVM.(RVP)}. Class PrimitivesSane(p: PrimitivesParams RiscvMachine): Prop := { getRegister_sane: forall r, mcomp_sane (getRegister r); @@ -92,6 +92,7 @@ Section Primitives. getPrivMode_sane: mcomp_sane getPrivMode; setPrivMode_sane: forall m, mcomp_sane (setPrivMode m); fence_sane: forall a b, mcomp_sane (fence a b); + leakEvent_sane: forall e, mcomp_sane (leakEvent e); getPC_sane: mcomp_sane getPC; setPC_sane: forall newPc, mcomp_sane (setPC newPc); endCycleNormal_sane: mcomp_sane endCycleNormal; @@ -151,15 +152,15 @@ Section Primitives. x = Register0 /\ post tt initialL) -> mcomp_sat (setRegister x v) initialL post; - spec_loadByte: spec_load 1 (Machine.loadByte (RiscvProgram := RVM)) Memory.loadByte; - spec_loadHalf: spec_load 2 (Machine.loadHalf (RiscvProgram := RVM)) Memory.loadHalf; - spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM)) Memory.loadWord; - spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM)) Memory.loadDouble; + spec_loadByte: spec_load 1 (Machine.loadByte (RiscvProgram := RVM.(RVP))) Memory.loadByte; + spec_loadHalf: spec_load 2 (Machine.loadHalf (RiscvProgram := RVM.(RVP))) Memory.loadHalf; + spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM.(RVP))) Memory.loadWord; + spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM.(RVP))) Memory.loadDouble; - spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM)) Memory.storeByte; - spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM)) Memory.storeHalf; - spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM)) Memory.storeWord; - spec_storeDouble: spec_store 8 (Machine.storeDouble (RiscvProgram := RVM)) Memory.storeDouble; + spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM.(RVP))) Memory.storeByte; + spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM.(RVP))) Memory.storeHalf; + spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM.(RVP))) Memory.storeWord; + spec_storeDouble: spec_store 8 (Machine.storeDouble (RiscvProgram := RVM.(RVP))) Memory.storeDouble; spec_getPC: forall initialL (post: word -> RiscvMachine -> Prop), post initialL.(getPc) initialL -> From 0218c3c5da186c64c586d91ec55e0546c3f4bb80 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 5 Nov 2023 20:59:49 -0500 Subject: [PATCH 11/42] wow it compiles --- src/riscv/Platform/MetricMinimal.v | 7 ++++++- src/riscv/Platform/MetricSane.v | 2 +- src/riscv/Platform/MinimalCSRs.v | 1 + src/riscv/Platform/MinimalMMIO_Post.v | 5 +++++ src/riscv/Platform/MinimalNoMul.v | 2 +- 5 files changed, 14 insertions(+), 3 deletions(-) diff --git a/src/riscv/Platform/MetricMinimal.v b/src/riscv/Platform/MetricMinimal.v index 6c1b83b..d7479ac 100644 --- a/src/riscv/Platform/MetricMinimal.v +++ b/src/riscv/Platform/MetricMinimal.v @@ -64,7 +64,12 @@ Section Riscv. fence := liftL2 id fence; endCycleNormal := liftL0 (addMetricInstructions 1) endCycleNormal; endCycleEarly{A} := liftL0 (addMetricInstructions 1) (@endCycleEarly _ _ _ _ _ A); - }. + }. + + Instance IsMetricRiscvMachineWithLeakage: @RiscvProgramWithLeakage (OState MetricRiscvMachine) _ _ _ := { + RVP := IsMetricRiscvMachine; + leakEvent := liftL1 id leakEvent; + }. Arguments Memory.load_bytes: simpl never. Arguments Memory.store_bytes: simpl never. diff --git a/src/riscv/Platform/MetricSane.v b/src/riscv/Platform/MetricSane.v index 5e5fde1..ccc6310 100644 --- a/src/riscv/Platform/MetricSane.v +++ b/src/riscv/Platform/MetricSane.v @@ -174,7 +174,7 @@ Section Sane. apply Bind_sane; [apply getPC_sane|intros]. apply Bind_sane; [apply loadWord_sane|intros]. apply Bind_sane; [apply leakage_of_instr_sane|intros]. - apply Bind_sane; [apply leakInstr_sane|intros]. + apply Bind_sane; [apply leakEvent_sane|intros]. apply Bind_sane; [apply execute_sane|intros]. apply endCycleNormal_sane. Qed. diff --git a/src/riscv/Platform/MinimalCSRs.v b/src/riscv/Platform/MinimalCSRs.v index 4340e55..1f61f08 100644 --- a/src/riscv/Platform/MinimalCSRs.v +++ b/src/riscv/Platform/MinimalCSRs.v @@ -99,6 +99,7 @@ Section Riscv. | Machine => postF tt mach | User | Supervisor => False end + | LeakEvent _ | MakeReservation _ | ClearReservation _ | CheckReservation _ diff --git a/src/riscv/Platform/MinimalMMIO_Post.v b/src/riscv/Platform/MinimalMMIO_Post.v index c6599c5..8e1510c 100644 --- a/src/riscv/Platform/MinimalMMIO_Post.v +++ b/src/riscv/Platform/MinimalMMIO_Post.v @@ -115,6 +115,11 @@ Section Riscv. - exact (post tt (withPc mach.(getNextPc) (withNextPc (word.add mach.(getNextPc) (word.of_Z 4)) mach))). Defined. + Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage (Post RiscvMachine) _ _ _ := + {| + RVP := IsRiscvMachine; + leakEvent _ := fun _ _ => False |}. + Definition MinimalMMIOPrimitivesParams: PrimitivesParams (Post RiscvMachine) RiscvMachine := {| Primitives.mcomp_sat A (m: Post RiscvMachine A) mach post := m mach post; Primitives.is_initial_register_value x := True; diff --git a/src/riscv/Platform/MinimalNoMul.v b/src/riscv/Platform/MinimalNoMul.v index d357e33..f588c0d 100644 --- a/src/riscv/Platform/MinimalNoMul.v +++ b/src/riscv/Platform/MinimalNoMul.v @@ -84,7 +84,7 @@ Section Riscv. postF tt (withNextPc (word.add mach.(getPc) (word.of_Z 4)) mach) | EndCycleNormal => fun postF postA => postF tt (updatePc mach) | EndCycleEarly _ => fun postF postA => postA (updatePc mach) (* ignores postF containing the continuation *) - | LeakEvent _ + | LeakEvent e => fun postF postA => postF tt (withLeakageEvent e mach) | MakeReservation _ | ClearReservation _ | CheckReservation _ From e81b8a1e33a30396337ad85f94f9833d3980624e Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 5 Nov 2023 21:07:18 -0500 Subject: [PATCH 12/42] turns out I was unable to build the compiler not because I needed to fix errors in riscv-coq, but rather because I needed to run make in the top-level directory. nice. From 0e6de80e04e6ccd12f80705ace3a7b05489796fb Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 5 Nov 2023 22:32:08 -0500 Subject: [PATCH 13/42] made leakage_of_instr a bit nicer to work with --- src/riscv/Platform/MetricSane.v | 2 +- src/riscv/Platform/Run.v | 2 +- src/riscv/Spec/Decode.v | 23 +++++++++++++++++++---- 3 files changed, 21 insertions(+), 6 deletions(-) diff --git a/src/riscv/Platform/MetricSane.v b/src/riscv/Platform/MetricSane.v index ccc6310..ba7b9d4 100644 --- a/src/riscv/Platform/MetricSane.v +++ b/src/riscv/Platform/MetricSane.v @@ -152,7 +152,7 @@ Section Sane. Ltac t := repeat (simpl; unfold when; repeat t_step'). Lemma leakage_of_instr_sane: forall inst, - mcomp_sane (leakage_of_instr getRegister inst). + mcomp_sane (leakage_of_instr regToZ_unsigned regToZ_signed getRegister inst). Proof. intros. destruct inst as [inst | inst | inst | inst | inst | inst | inst | inst | inst | inst]; diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index ea86151..808ec92 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -25,7 +25,7 @@ Section Riscv. pc <- getPC; inst <- loadWord Fetch pc; let inst' := decode iset (combine 4 inst) in - leakage_event <- leakage_of_instr getRegister inst'; + leakage_event <- leakage_of_instr regToZ_unsigned regToZ_signed getRegister inst'; leakEvent leakage_event;; execute inst';; endCycleNormal. diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 8e96075..275b52e 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -31,8 +31,8 @@ Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Section WithMonad. Context {M : Type -> Type} {MM : Monad M}. - Context {mword: Type} {MW: MachineWidth mword}. - Context (getRegister : Register -> M mword). + Context {mword: Type}. + Context (regToZ_unsigned: mword -> Z) (regToZ_signed: mword -> Z) (getRegister : Register -> M mword). (* Converted type declarations: *) @@ -152,8 +152,6 @@ Inductive LeakageI64 : Type := | Sraw_leakage | InvalidI64_leakage. Locate "<- ; _ ". -Print MW. Check add. Check regToZ_unsigned. - Definition leakage_of_instr_I64 (instr : InstructionI64) : M LeakageI64 := match instr with | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (regToZ_unsigned rs1_val + oimm12))) @@ -1909,3 +1907,20 @@ Definition decode Utility.Utility.bitSlice Utility.Utility.machineIntToShamt Utility.Utility.signExtend *) + +Definition Mtriv (x : Type) := x. +Definition trivialBind (A B : Type) (x : Mtriv A) (f : A -> B) : Mtriv B := f x. +Definition trivialReturn (A : Type) (a : A) : Mtriv A := a. +Print Monad. +Lemma trivial_left_identity : forall (A B : Type) (a : A) (f : A -> Mtriv B), trivialBind A B (trivialReturn A a) f = f a. +Proof. trivial. Qed. +Lemma trivial_right_identity : forall (A : Type) (m : Mtriv A), trivialBind A A m (trivialReturn A) = m. +Proof. trivial. Qed. +Lemma trivial_associativity : forall (A B C : Type) (m : Mtriv A) (f : A -> Mtriv B) (g : B -> Mtriv C), trivialBind B C (trivialBind A B m f) g = trivialBind A C m (fun x : A => trivialBind B C (f x) g). +Proof. trivial. Qed. +Check @leakage_of_instr. +Definition trivialMonad : Monad Mtriv := + {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. +Definition concrete_leakage_of_instr := @leakage_of_instr Mtriv trivialMonad. +Definition concrete_leakage_of_instr_Z := concrete_leakage_of_instr Z (fun x => x) (fun x => x). +Print concrete_leakage_of_instr_Z. Print leakage_of_instr. From 0d78de0a8eeb255eaf0ab2a2f3428043ae25f520 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sat, 18 Nov 2023 00:35:06 -0500 Subject: [PATCH 14/42] changed leakage trace to use word instead of Z --- src/riscv/Spec/Decode.v | 169 ++++++++++++++++++++-------------------- 1 file changed, 84 insertions(+), 85 deletions(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 275b52e..218bd72 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -31,8 +31,9 @@ Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Section WithMonad. Context {M : Type -> Type} {MM : Monad M}. - Context {mword: Type}. - Context (regToZ_unsigned: mword -> Z) (regToZ_signed: mword -> Z) (getRegister : Register -> M mword). + Context {width} {BW : Bitwidth width} {word: word.word width} {word_ok: word.ok word}. + Context (getRegister : Register -> M word). + (*Am I handling signed/unsigned stuff correctly?*) (* Converted type declarations: *) @@ -138,29 +139,29 @@ Inductive InstructionI64 : Type := | InvalidI64 : InstructionI64. Inductive LeakageI64 : Type := -| Ld_leakage (addr: Utility.Utility.MachineInt) -| Lwu_leakage (addr: Utility.Utility.MachineInt) +| Ld_leakage (addr: word) +| Lwu_leakage (addr: word) | Addiw_leakage | Slliw_leakage | Srliw_leakage | Sraiw_leakage -| Sd_leakage (addr: Utility.Utility.MachineInt) +| Sd_leakage (addr: word) | Addw_leakage | Subw_leakage | Sllw_leakage | Srlw_leakage | Sraw_leakage -| InvalidI64_leakage. Locate "<- ; _ ". +| InvalidI64_leakage. Definition leakage_of_instr_I64 (instr : InstructionI64) : M LeakageI64 := match instr with - | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (regToZ_unsigned rs1_val + oimm12))) - | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (regToZ_unsigned rs1_val + oimm12))) + | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (word.add rs1_val (word.of_Z oimm12)))) + | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (word.add rs1_val (word.of_Z oimm12)))) | Addiw _ _ _ => Return Addiw_leakage | Slliw _ _ _ => Return Slliw_leakage | Srliw _ _ _ => Return Srliw_leakage | Sraiw _ _ _ => Return Sraiw_leakage - | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (regToZ_unsigned rs1_val + simm12))) + | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (word.add rs1_val (word.of_Z simm12)))) | Addw _ _ _ => Return Addw_leakage | Subw _ _ _ => Return Subw_leakage | Sllw _ _ _ => Return Sllw_leakage @@ -170,11 +171,11 @@ Definition leakage_of_instr_I64 (instr : InstructionI64) : M LeakageI64 := end. Inductive LeakageI : Type := -| Lb_leakage (addr: Utility.Utility.MachineInt) -| Lh_leakage (addr: Utility.Utility.MachineInt) -| Lw_leakage (addr: Utility.Utility.MachineInt) -| Lbu_leakage (addr: Utility.Utility.MachineInt) -| Lhu_leakage (addr: Utility.Utility.MachineInt) +| Lb_leakage (addr: word) +| Lh_leakage (addr: word) +| Lw_leakage (addr: word) +| Lbu_leakage (addr: word) +| Lhu_leakage (addr: word) | Fence_leakage (* unsure about this one. *) | Fence_i_leakage | Addi_leakage @@ -187,9 +188,9 @@ Inductive LeakageI : Type := | Srli_leakage | Srai_leakage | Auipc_leakage -| Sb_leakage (addr: Utility.Utility.MachineInt) -| Sh_leakage (addr: Utility.Utility.MachineInt) -| Sw_leakage (addr: Utility.Utility.MachineInt) +| Sb_leakage (addr: word) +| Sh_leakage (addr: word) +| Sw_leakage (addr: word) | Add_leakage | Sub_leakage | Sll_leakage @@ -275,14 +276,14 @@ Inductive InstructionI : Type := | Jal (rd : Register) (jimm20 : Utility.Utility.MachineInt) : InstructionI | InvalidI : InstructionI. Search Z. Compute (Utility.regToZ_unsigned 0). -(* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) +(* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) Print word.word. Definition leakage_of_instr_I (instr : InstructionI) : M LeakageI := match instr with - | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (regToZ_unsigned rs1_val + oimm12)) - | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (regToZ_unsigned rs1_val + oimm12)) - | Lw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lw_leakage (regToZ_unsigned rs1_val + oimm12)) - | Lbu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lbu_leakage (regToZ_unsigned rs1_val + oimm12)) - | Lhu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lhu_leakage (regToZ_unsigned rs1_val + oimm12)) + | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (word.add rs1_val (word.of_Z oimm12))) + | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (word.add rs1_val (word.of_Z oimm12))) + | Lw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lw_leakage (word.add rs1_val (word.of_Z oimm12))) + | Lbu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lbu_leakage (word.add rs1_val (word.of_Z oimm12))) + | Lhu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lhu_leakage (word.add rs1_val (word.of_Z oimm12))) | Fence _ _ => Return Fence_leakage | Fence_i => Return Fence_i_leakage | Addi _ _ _ => Return Addi_leakage @@ -295,9 +296,9 @@ Definition leakage_of_instr_I (instr : InstructionI) : M LeakageI := | Srli _ _ _ => Return Srli_leakage | Srai _ _ _ => Return Srai_leakage | Auipc _ _ => Return Auipc_leakage - | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (regToZ_unsigned rs1_val + simm12)) - | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (regToZ_unsigned rs1_val + simm12)) - | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (regToZ_unsigned rs1_val + simm12)) + | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (word.add rs1_val (word.of_Z simm12))) + | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (word.add rs1_val (word.of_Z simm12))) + | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (word.add rs1_val (word.of_Z simm12))) | Add _ _ _ => Return Add_leakage | Sub _ _ _ => Return Sub_leakage | Sll _ _ _ => Return Sll_leakage @@ -309,12 +310,12 @@ Definition leakage_of_instr_I (instr : InstructionI) : M LeakageI := | Or _ _ _ => Return Or_leakage | And _ _ _ => Return And_leakage | Lui _ _ => Return Lui_leakage - | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Beq_leakage (regToZ_unsigned rs1_val =? regToZ_unsigned rs2_val)) - | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bne_leakage (negb (regToZ_unsigned rs1_val =? regToZ_unsigned rs2_val))) - | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Blt_leakage (regToZ_signed rs1_val rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bge_leakage (regToZ_signed rs1_val >=? regToZ_signed rs2_val)) - | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bltu_leakage (regToZ_unsigned rs1_val <=? regToZ_unsigned rs2_val)) - | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bgeu_leakage (regToZ_unsigned rs1_val >=? regToZ_unsigned rs2_val)) + | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Beq_leakage (word.eqb rs1_val rs2_val)) + | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bne_leakage (negb (word.eqb rs1_val rs2_val))) + | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Blt_leakage (word.lts rs1_val rs2_val)) + | Bge rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bge_leakage (negb (word.lts rs1_val rs2_val))) + | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bltu_leakage (word.ltu rs1_val rs2_val)) + | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bgeu_leakage (negb (word.ltu rs1_val rs2_val))) | Jalr _ _ _ => Return Jalr_leakage | Jal _ _ => Return Jal_leakage | InvalidI => Return InvalidI_leakage @@ -344,8 +345,8 @@ Definition leakage_of_instr_F64 (instr : InstructionF64) : M LeakageF64 := end. Inductive LeakageF : Type := -| Flw_leakage (addr: Utility.Utility.MachineInt) -| Fsw_leakage (addr: Utility.Utility.MachineInt) +| Flw_leakage (addr: word) +| Fsw_leakage (addr: word) | Fmadd_s_leakage | Fmsub_s_leakage | Fnmsub_s_leakage @@ -419,12 +420,12 @@ Inductive InstructionF : Type := | Fcvt_s_w (rd : FPRegister) (rs1 : Register) (rm : RoundMode) : InstructionF | Fcvt_s_wu (rd : FPRegister) (rs1 : Register) (rm : RoundMode) : InstructionF | Fmv_w_x (rd : FPRegister) (rs1 : Register) : InstructionF - | InvalidF : InstructionF. + | InvalidF : InstructionF. Print word.word. Definition leakage_of_instr_F (instr : InstructionF) : M LeakageF := match instr with - | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (regToZ_unsigned rs1_val + oimm12)) - | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (regToZ_unsigned rs1_val + simm12)) + | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (word.add rs1_val (word.of_Z oimm12))) + | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (word.add rs1_val (word.of_Z simm12))) | Fmadd_s _ _ _ _ _ => Return Fmadd_s_leakage | Fmsub_s _ _ _ _ _ => Return Fmsub_s_leakage | Fnmsub_s _ _ _ _ _ => Return Fnmsub_s_leakage @@ -515,17 +516,17 @@ Definition leakage_of_instr_CSR (instr : InstructionCSR) : M LeakageCSR := (* do we care about aqrl here? *) Inductive LeakageA64 : Type := -| Lr_d_leakage (addr : Utility.Utility.MachineInt) -| Sc_d_leakage (addr : Utility.Utility.MachineInt) (* behavior of this depends on whether there is a reservation on addr... *) -| Amoswap_d_leakage (addr : Utility.Utility.MachineInt) -| Amoadd_d_leakage (addr : Utility.Utility.MachineInt) -| Amoand_d_leakage (addr : Utility.Utility.MachineInt) -| Amoor_d_leakage (addr : Utility.Utility.MachineInt) -| Amoxor_d_leakage (addr : Utility.Utility.MachineInt) -| Amomax_d_leakage (addr : Utility.Utility.MachineInt) -| Amomaxu_d_leakage (addr : Utility.Utility.MachineInt) -| Amomin_d_leakage (addr : Utility.Utility.MachineInt) -| Amominu_d_leakage (addr : Utility.Utility.MachineInt) +| Lr_d_leakage (addr : word) +| Sc_d_leakage (addr : word) (* behavior of this depends on whether there is a reservation on addr... *) +| Amoswap_d_leakage (addr : word) +| Amoadd_d_leakage (addr : word) +| Amoand_d_leakage (addr : word) +| Amoor_d_leakage (addr : word) +| Amoxor_d_leakage (addr : word) +| Amomax_d_leakage (addr : word) +| Amomaxu_d_leakage (addr : word) +| Amomin_d_leakage (addr : word) +| Amominu_d_leakage (addr : word) | InvalidA64_leakage. Inductive InstructionA64 : Type := @@ -565,32 +566,32 @@ Inductive InstructionA64 : Type := Definition leakage_of_instr_A64 (instr : InstructionA64) : M LeakageA64 := match instr with - | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage (regToZ_unsigned rs1_val)) - | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage (regToZ_unsigned rs1_val)) - | Amoswap_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_d_leakage (regToZ_unsigned rs1_val)) - | Amoadd_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_d_leakage (regToZ_unsigned rs1_val)) - | Amoand_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_d_leakage (regToZ_unsigned rs1_val)) - | Amoor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_d_leakage (regToZ_unsigned rs1_val)) - | Amoxor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_d_leakage (regToZ_unsigned rs1_val)) - | Amomax_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_d_leakage (regToZ_unsigned rs1_val)) - | Amomaxu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_d_leakage (regToZ_unsigned rs1_val)) - | Amomin_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_d_leakage (regToZ_unsigned rs1_val)) - | Amominu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_d_leakage (regToZ_unsigned rs1_val)) + | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage rs1_val) + | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage rs1_val) + | Amoswap_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_d_leakage rs1_val) + | Amoadd_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_d_leakage rs1_val) + | Amoand_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_d_leakage rs1_val) + | Amoor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_d_leakage rs1_val) + | Amoxor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_d_leakage rs1_val) + | Amomax_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_d_leakage rs1_val) + | Amomaxu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_d_leakage rs1_val) + | Amomin_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_d_leakage rs1_val) + | Amominu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_d_leakage rs1_val) | InvalidA64 => Return InvalidA64_leakage end. Inductive LeakageA : Type := -| Lr_w_leakage (addr : Utility.Utility.MachineInt) -| Sc_w_leakage (addr : Utility.Utility.MachineInt) -| Amoswap_w_leakage (addr : Utility.Utility.MachineInt) -| Amoadd_w_leakage (addr : Utility.Utility.MachineInt) -| Amoand_w_leakage (addr : Utility.Utility.MachineInt) -| Amoor_w_leakage (addr : Utility.Utility.MachineInt) -| Amoxor_w_leakage (addr : Utility.Utility.MachineInt) -| Amomax_w_leakage (addr : Utility.Utility.MachineInt) -| Amomaxu_w_leakage (addr : Utility.Utility.MachineInt) -| Amomin_w_leakage (addr : Utility.Utility.MachineInt) -| Amominu_w_leakage (addr : Utility.Utility.MachineInt) +| Lr_w_leakage (addr : word) +| Sc_w_leakage (addr : word) +| Amoswap_w_leakage (addr : word) +| Amoadd_w_leakage (addr : word) +| Amoand_w_leakage (addr : word) +| Amoor_w_leakage (addr : word) +| Amoxor_w_leakage (addr : word) +| Amomax_w_leakage (addr : word) +| Amomaxu_w_leakage (addr : word) +| Amomin_w_leakage (addr : word) +| Amominu_w_leakage (addr : word) | InvalidA_leakage. Inductive InstructionA : Type := @@ -630,17 +631,17 @@ Inductive InstructionA : Type := Definition leakage_of_instr_A (instr : InstructionA) : M LeakageA := match instr with - | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage (regToZ_unsigned rs1_val)) - | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage (regToZ_unsigned rs1_val)) - | Amoswap_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_w_leakage (regToZ_unsigned rs1_val)) - | Amoadd_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_w_leakage (regToZ_unsigned rs1_val)) - | Amoand_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_w_leakage (regToZ_unsigned rs1_val)) - | Amoor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_w_leakage (regToZ_unsigned rs1_val)) - | Amoxor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_w_leakage (regToZ_unsigned rs1_val)) - | Amomax_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_w_leakage (regToZ_unsigned rs1_val)) - | Amomaxu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_w_leakage (regToZ_unsigned rs1_val)) - | Amomin_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_w_leakage (regToZ_unsigned rs1_val)) - | Amominu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_w_leakage (regToZ_unsigned rs1_val)) + | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage rs1_val) + | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage rs1_val) + | Amoswap_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_w_leakage rs1_val) + | Amoadd_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_w_leakage rs1_val) + | Amoand_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_w_leakage rs1_val) + | Amoor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_w_leakage rs1_val) + | Amoxor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_w_leakage rs1_val) + | Amomax_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_w_leakage rs1_val) + | Amomaxu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_w_leakage rs1_val) + | Amomin_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_w_leakage rs1_val) + | Amominu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_w_leakage rs1_val) | InvalidA => Return InvalidA_leakage end. @@ -666,7 +667,7 @@ Inductive LeakageEvent : Type := | A64Leakage (a64Leakage : LeakageA64) | F64Leakage (f64Leakage : LeakageF64) | CSRLeakage (csrLeakage : LeakageCSR) -| InvalidLeakage. Check @Bind. +| InvalidLeakage. Definition leakage_of_instr (instr : Instruction) : M LeakageEvent := match instr with @@ -1921,6 +1922,4 @@ Proof. trivial. Qed. Check @leakage_of_instr. Definition trivialMonad : Monad Mtriv := {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. -Definition concrete_leakage_of_instr := @leakage_of_instr Mtriv trivialMonad. -Definition concrete_leakage_of_instr_Z := concrete_leakage_of_instr Z (fun x => x) (fun x => x). -Print concrete_leakage_of_instr_Z. Print leakage_of_instr. +Definition concrete_leakage_of_instr := @leakage_of_instr Mtriv trivialMonad. Check concrete_leakage_of_instr. From 9dfe76e6fc043255575745277bdb8507e1f65209 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sat, 18 Nov 2023 01:09:33 -0500 Subject: [PATCH 15/42] started making things build again. not done. --- src/riscv/Platform/MaterializeRiscvProgram.v | 2 +- src/riscv/Platform/MetricMinimal.v | 2 +- src/riscv/Platform/MetricSane.v | 2 +- src/riscv/Platform/Minimal.v | 2 +- src/riscv/Platform/MinimalMMIO_Post.v | 2 +- src/riscv/Platform/RiscvMachine.v | 2 +- src/riscv/Platform/Run.v | 20 ++++++++++---------- src/riscv/Spec/Decode.v | 2 +- src/riscv/Spec/Machine.v | 8 +++++--- 9 files changed, 22 insertions(+), 20 deletions(-) diff --git a/src/riscv/Platform/MaterializeRiscvProgram.v b/src/riscv/Platform/MaterializeRiscvProgram.v index 36de0e1..d80d78d 100644 --- a/src/riscv/Platform/MaterializeRiscvProgram.v +++ b/src/riscv/Platform/MaterializeRiscvProgram.v @@ -27,7 +27,7 @@ Section Riscv. | GetPrivMode | SetPrivMode (_ : PrivMode) | Fence (_ : MachineInt) (_ : MachineInt) - | LeakEvent (_ : LeakageEvent) + | LeakEvent (_ : @LeakageEvent width word) | GetPC | SetPC (_ : word) | StartCycle diff --git a/src/riscv/Platform/MetricMinimal.v b/src/riscv/Platform/MetricMinimal.v index d7479ac..e8fcd9a 100644 --- a/src/riscv/Platform/MetricMinimal.v +++ b/src/riscv/Platform/MetricMinimal.v @@ -66,7 +66,7 @@ Section Riscv. endCycleEarly{A} := liftL0 (addMetricInstructions 1) (@endCycleEarly _ _ _ _ _ A); }. - Instance IsMetricRiscvMachineWithLeakage: @RiscvProgramWithLeakage (OState MetricRiscvMachine) _ _ _ := { + Instance IsMetricRiscvMachineWithLeakage: @RiscvProgramWithLeakage width BW word (OState MetricRiscvMachine) _ _ _ := { RVP := IsMetricRiscvMachine; leakEvent := liftL1 id leakEvent; }. diff --git a/src/riscv/Platform/MetricSane.v b/src/riscv/Platform/MetricSane.v index ba7b9d4..ccc6310 100644 --- a/src/riscv/Platform/MetricSane.v +++ b/src/riscv/Platform/MetricSane.v @@ -152,7 +152,7 @@ Section Sane. Ltac t := repeat (simpl; unfold when; repeat t_step'). Lemma leakage_of_instr_sane: forall inst, - mcomp_sane (leakage_of_instr regToZ_unsigned regToZ_signed getRegister inst). + mcomp_sane (leakage_of_instr getRegister inst). Proof. intros. destruct inst as [inst | inst | inst | inst | inst | inst | inst | inst | inst | inst]; diff --git a/src/riscv/Platform/Minimal.v b/src/riscv/Platform/Minimal.v index 5f49695..c6f39b7 100644 --- a/src/riscv/Platform/Minimal.v +++ b/src/riscv/Platform/Minimal.v @@ -101,7 +101,7 @@ Section Riscv. }. Print RiscvProgramWithLeakage. - Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage (OState RiscvMachine) _ _ _ := { + Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage width BW word (OState RiscvMachine) _ _ _ := { RVP := IsRiscvMachine; leakEvent e := fail_hard; }. diff --git a/src/riscv/Platform/MinimalMMIO_Post.v b/src/riscv/Platform/MinimalMMIO_Post.v index 8e1510c..9d1a8f3 100644 --- a/src/riscv/Platform/MinimalMMIO_Post.v +++ b/src/riscv/Platform/MinimalMMIO_Post.v @@ -115,7 +115,7 @@ Section Riscv. - exact (post tt (withPc mach.(getNextPc) (withNextPc (word.add mach.(getNextPc) (word.of_Z 4)) mach))). Defined. - Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage (Post RiscvMachine) _ _ _ := + Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage width BW word (Post RiscvMachine) _ _ _ := {| RVP := IsRiscvMachine; leakEvent _ := fun _ _ => False |}. diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index da30e7e..6aaf2a2 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -141,7 +141,7 @@ Section Machine. getMem: Mem; getXAddrs: XAddrs; getLog: list LogItem; - getTrace: list LeakageEvent; + getTrace: list (@LeakageEvent width word); }. Definition withRegs: Registers -> RiscvMachine -> RiscvMachine := diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index 808ec92..d230bb6 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -10,25 +10,25 @@ Require Import riscv.Utility.Utility. Section Riscv. - Context {mword: Type}. + Context {width} {BW : Bitwidth width} {mword: word.word width}. Context {MW: MachineWidth mword}. Context {M: Type -> Type}. Context {MM: Monad M}. Context {RVM: RiscvProgramWithLeakage}. - Context {RVS: RiscvMachine M mword}. - - Print leakage_of_instr. Print decode. Print leakEvent. Print getRegister. Print MachineWidth. + Context {RVS: RiscvMachine M mword}. Print leakage_of_instr. + Check (leakage_of_instr getRegister _). + Print leakage_of_instr. Check getRegister. Definition run1(iset: InstructionSet): M unit := pc <- getPC; - inst <- loadWord Fetch pc; - let inst' := decode iset (combine 4 inst) in - leakage_event <- leakage_of_instr regToZ_unsigned regToZ_signed getRegister inst'; - leakEvent leakage_event;; - execute inst';; - endCycleNormal. + inst <- loadWord Fetch pc; + let inst' := decode iset (combine 4 inst) in + leakage_event <- leakage_of_instr getRegister inst'; + leakEvent leakage_event;; + execute inst';; + endCycleNormal. (* Note: We cannot use power_func (fun m => run1 iset;; m) n (Return tt) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 218bd72..e4bce62 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -31,7 +31,7 @@ Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Section WithMonad. Context {M : Type -> Type} {MM : Monad M}. - Context {width} {BW : Bitwidth width} {word: word.word width} {word_ok: word.ok word}. + Context {width} {BW : Bitwidth width} {word: word.word width}. Context (getRegister : Register -> M word). (*Am I handling signed/unsigned stuff correctly?*) diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index b814426..e473b1e 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -64,12 +64,14 @@ Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { Some instances may not support endCycleEarly and fail instead. *) endCycleNormal: M unit; endCycleEarly: forall A, M A; -}. +}. Print LeakageEvent. -Class RiscvProgramWithLeakage{M}{t}{MM}{MWt}:= +Class RiscvProgramWithLeakage + {width} {BW : Bitwidth width} {word: word.word width} + {M}{t}{MM}{MWt}:= mkRiscvProgramWithLeakage { RVP :> @RiscvProgram M t MM MWt; - leakEvent : LeakageEvent -> M unit; + leakEvent : @LeakageEvent width word -> M unit; }. Class RiscvMachine`{MP: RiscvProgram} := mkRiscvMachine { From 996df7914f920b19e44b9d2abccc11ff4e0bfb96 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sat, 18 Nov 2023 04:07:14 -0500 Subject: [PATCH 16/42] made things be inferred more easily --- src/riscv/Platform/MaterializeRiscvProgram.v | 2 +- src/riscv/Platform/RiscvMachine.v | 2 +- src/riscv/Spec/Decode.v | 61 +++++++++++++------- src/riscv/Spec/Machine.v | 2 +- 4 files changed, 44 insertions(+), 23 deletions(-) diff --git a/src/riscv/Platform/MaterializeRiscvProgram.v b/src/riscv/Platform/MaterializeRiscvProgram.v index d80d78d..36de0e1 100644 --- a/src/riscv/Platform/MaterializeRiscvProgram.v +++ b/src/riscv/Platform/MaterializeRiscvProgram.v @@ -27,7 +27,7 @@ Section Riscv. | GetPrivMode | SetPrivMode (_ : PrivMode) | Fence (_ : MachineInt) (_ : MachineInt) - | LeakEvent (_ : @LeakageEvent width word) + | LeakEvent (_ : LeakageEvent) | GetPC | SetPC (_ : word) | StartCycle diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index 6aaf2a2..da30e7e 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -141,7 +141,7 @@ Section Machine. getMem: Mem; getXAddrs: XAddrs; getLog: list LogItem; - getTrace: list (@LeakageEvent width word); + getTrace: list LeakageEvent; }. Definition withRegs: Registers -> RiscvMachine -> RiscvMachine := diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index e4bce62..1d617c1 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -31,9 +31,6 @@ Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Section WithMonad. Context {M : Type -> Type} {MM : Monad M}. - Context {width} {BW : Bitwidth width} {word: word.word width}. - Context (getRegister : Register -> M word). - (*Am I handling signed/unsigned stuff correctly?*) (* Converted type declarations: *) @@ -138,7 +135,9 @@ Inductive InstructionI64 : Type := | Sraw (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionI64 | InvalidI64 : InstructionI64. -Inductive LeakageI64 : Type := +Inductive LeakageI64 +{width} {BW : Bitwidth width} {word: word.word width} + : Type := | Ld_leakage (addr: word) | Lwu_leakage (addr: word) | Addiw_leakage @@ -153,7 +152,9 @@ Inductive LeakageI64 : Type := | Sraw_leakage | InvalidI64_leakage. -Definition leakage_of_instr_I64 (instr : InstructionI64) : M LeakageI64 := +Definition leakage_of_instr_I64 + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionI64) : M LeakageI64 := match instr with | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (word.add rs1_val (word.of_Z oimm12)))) | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (word.add rs1_val (word.of_Z oimm12)))) @@ -170,7 +171,9 @@ Definition leakage_of_instr_I64 (instr : InstructionI64) : M LeakageI64 := | InvalidI64 => Return InvalidI64_leakage end. -Inductive LeakageI : Type := +Inductive LeakageI + {width} {BW : Bitwidth width} {word: word.word width} + : Type := | Lb_leakage (addr: word) | Lh_leakage (addr: word) | Lw_leakage (addr: word) @@ -277,7 +280,9 @@ Inductive InstructionI : Type := | InvalidI : InstructionI. Search Z. Compute (Utility.regToZ_unsigned 0). (* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) Print word.word. -Definition leakage_of_instr_I (instr : InstructionI) : M LeakageI := +Definition leakage_of_instr_I + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionI) : M LeakageI := match instr with | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (word.add rs1_val (word.of_Z oimm12))) | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (word.add rs1_val (word.of_Z oimm12))) @@ -344,7 +349,9 @@ Definition leakage_of_instr_F64 (instr : InstructionF64) : M LeakageF64 := | InvalidF64 => Return InvalidF64_leakage end. -Inductive LeakageF : Type := +Inductive LeakageF + {width} {BW : Bitwidth width} {word: word.word width} + : Type := | Flw_leakage (addr: word) | Fsw_leakage (addr: word) | Fmadd_s_leakage @@ -422,7 +429,9 @@ Inductive InstructionF : Type := | Fmv_w_x (rd : FPRegister) (rs1 : Register) : InstructionF | InvalidF : InstructionF. Print word.word. -Definition leakage_of_instr_F (instr : InstructionF) : M LeakageF := +Definition leakage_of_instr_F + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionF) : M LeakageF := match instr with | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (word.add rs1_val (word.of_Z oimm12))) | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (word.add rs1_val (word.of_Z simm12))) @@ -515,7 +524,9 @@ Definition leakage_of_instr_CSR (instr : InstructionCSR) : M LeakageCSR := end. (* do we care about aqrl here? *) -Inductive LeakageA64 : Type := +Inductive LeakageA64 + {width} {BW : Bitwidth width} {word: word.word width} + : Type := | Lr_d_leakage (addr : word) | Sc_d_leakage (addr : word) (* behavior of this depends on whether there is a reservation on addr... *) | Amoswap_d_leakage (addr : word) @@ -564,7 +575,9 @@ Inductive InstructionA64 : Type := : InstructionA64 | InvalidA64 : InstructionA64. -Definition leakage_of_instr_A64 (instr : InstructionA64) : M LeakageA64 := +Definition leakage_of_instr_A64 + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionA64) : M LeakageA64 := match instr with | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage rs1_val) | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage rs1_val) @@ -580,7 +593,9 @@ Definition leakage_of_instr_A64 (instr : InstructionA64) : M LeakageA64 := | InvalidA64 => Return InvalidA64_leakage end. -Inductive LeakageA : Type := + Inductive LeakageA + {width} {BW : Bitwidth width} {word: word.word width} + : Type := | Lr_w_leakage (addr : word) | Sc_w_leakage (addr : word) | Amoswap_w_leakage (addr : word) @@ -629,7 +644,9 @@ Inductive InstructionA : Type := : InstructionA | InvalidA : InstructionA. -Definition leakage_of_instr_A (instr : InstructionA) : M LeakageA := +Definition leakage_of_instr_A + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionA) : M LeakageA := match instr with | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage rs1_val) | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage rs1_val) @@ -657,7 +674,9 @@ Inductive Instruction : Type := | CSRInstruction (csrInstruction : InstructionCSR) : Instruction | InvalidInstruction (inst : Utility.Utility.MachineInt) : Instruction. -Inductive LeakageEvent : Type := +Inductive LeakageEvent + {width} {BW : Bitwidth width} {word: word.word width} + : Type := | ILeakage (iLeakage : LeakageI) | MLeakage (mLeakage : LeakageM) | ALeakage (aLeakage : LeakageA) @@ -669,15 +688,17 @@ Inductive LeakageEvent : Type := | CSRLeakage (csrLeakage : LeakageCSR) | InvalidLeakage. -Definition leakage_of_instr (instr : Instruction) : M LeakageEvent := +Definition leakage_of_instr + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : Instruction) : M LeakageEvent := match instr with - | IInstruction instr => l <- leakage_of_instr_I instr; Return (ILeakage l) + | IInstruction instr => l <- leakage_of_instr_I getRegister instr; Return (ILeakage l) | MInstruction instr => l <- leakage_of_instr_M instr; Return (MLeakage l) - | AInstruction instr => l <- leakage_of_instr_A instr; Return (ALeakage l) - | FInstruction instr => l <- leakage_of_instr_F instr; Return (FLeakage l) - | I64Instruction instr => l <- leakage_of_instr_I64 instr; Return (I64Leakage l) + | AInstruction instr => l <- leakage_of_instr_A getRegister instr; Return (ALeakage l) + | FInstruction instr => l <- leakage_of_instr_F getRegister instr; Return (FLeakage l) + | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; Return (I64Leakage l) | M64Instruction instr => l <- leakage_of_instr_M64 instr; Return (M64Leakage l) - | A64Instruction instr => l <- leakage_of_instr_A64 instr; Return (A64Leakage l) + | A64Instruction instr => l <- leakage_of_instr_A64 getRegister instr; Return (A64Leakage l) | F64Instruction instr => l <- leakage_of_instr_F64 instr; Return (F64Leakage l) | CSRInstruction instr => l <- leakage_of_instr_CSR instr; Return (CSRLeakage l) | InvalidInstruction _ => Return InvalidLeakage diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index e473b1e..db4d734 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -71,7 +71,7 @@ Class RiscvProgramWithLeakage {M}{t}{MM}{MWt}:= mkRiscvProgramWithLeakage { RVP :> @RiscvProgram M t MM MWt; - leakEvent : @LeakageEvent width word -> M unit; + leakEvent : @LeakageEvent width BW word -> M unit; }. Class RiscvMachine`{MP: RiscvProgram} := mkRiscvMachine { From b6f4d400bdfbca1766213ca69e156fb95fb59339 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sat, 18 Nov 2023 04:16:54 -0500 Subject: [PATCH 17/42] fix implicit args --- src/riscv/Spec/Decode.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 1d617c1..c637fd7 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -1943,4 +1943,4 @@ Proof. trivial. Qed. Check @leakage_of_instr. Definition trivialMonad : Monad Mtriv := {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. -Definition concrete_leakage_of_instr := @leakage_of_instr Mtriv trivialMonad. Check concrete_leakage_of_instr. +Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := @leakage_of_instr Mtriv trivialMonad. Check concrete_leakage_of_instr. From 7d71fee7de6298aae4503f15d124b6ab6976faa8 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sat, 18 Nov 2023 14:04:30 -0500 Subject: [PATCH 18/42] fixed implicit args --- src/riscv/Spec/Decode.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index c637fd7..973d918 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -1943,4 +1943,4 @@ Proof. trivial. Qed. Check @leakage_of_instr. Definition trivialMonad : Monad Mtriv := {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. -Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := @leakage_of_instr Mtriv trivialMonad. Check concrete_leakage_of_instr. +Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := @leakage_of_instr Mtriv trivialMonad width BW word. From 6f331f5f1a47b87a0c76498dfa88d77dadbd6264 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Tue, 12 Dec 2023 19:47:51 -0500 Subject: [PATCH 19/42] Changed assembly semantics as follows. shift time depends on shift amount. division is not constant time. --- src/riscv/Spec/Decode.v | 123 ++++++++++++++++++++-------------------- 1 file changed, 63 insertions(+), 60 deletions(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 973d918..de3028e 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -60,23 +60,24 @@ Inductive InstructionM64 : Type := | Remuw (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM64 | InvalidM64 : InstructionM64. -(* I believe division is often not implemented in a constant-time way, so this will have to - be changed at some point. *) -Inductive LeakageM64 : Type := -| Mulw_leakage -| Divw_leakage -| Divuw_leakage -| Remw_leakage -| Remuw_leakage +Inductive LeakageM64 + {width} {BW : Bitwidth width} {word: word.word width} : Type := +| Mulw_leakage (*andres says that this is constant-time*) +| Divw_leakage (num : word) (den : word) (*but that this is not.*) +| Divuw_leakage (num : word) (den : word) +| Remw_leakage (num : word) (den : word) (*i'm not sure about this one. probably same as div, so not constant-time.*) +| Remuw_leakage (num : word) (den : word) | InvalidM64_leakage. -Definition leakage_of_instr_M64 (instr : InstructionM64) : M LeakageM64 := +Definition leakage_of_instr_M64 + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionM64) : M LeakageM64 := match instr with | Mulw _ _ _ => Return Mulw_leakage - | Divw _ _ _ => Return Divw_leakage - | Divuw _ _ _ => Return Divw_leakage - | Remw _ _ _ => Return Remw_leakage - | Remuw _ _ _ => Return Remuw_leakage + | Divw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divw_leakage num den) + | Divuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divuw_leakage num den) + | Remw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remw_leakage num den) + | Remuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remuw_leakage num den) | InvalidM64 => Return InvalidM64_leakage end. @@ -91,28 +92,30 @@ Inductive InstructionM : Type := | Remu (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM | InvalidM : InstructionM. -(* This will also have to change, due to division beingn implemented in a not-constant-time way. *) -Inductive LeakageM : Type := +Inductive LeakageM + {width} {BW : Bitwidth width} {word: word.word width} : Type := | Mul_leakage | Mulh_leakage | Mulhsu_leakage | Mulhu_leakage -| Div_leakage -| Divu_leakage -| Rem_leakage -| Remu_leakage +| Div_leakage (num : word) (den : word) +| Divu_leakage (num : word) (den : word) +| Rem_leakage (num : word) (den : word) +| Remu_leakage (num : word) (den : word) | InvalidM_leakage. -Definition leakage_of_instr_M (instr : InstructionM) : M LeakageM := +Definition leakage_of_instr_M + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionM) : M LeakageM := match instr with | Mul _ _ _ => Return Mul_leakage | Mulh _ _ _ => Return Mulh_leakage | Mulhsu _ _ _ => Return Mulhsu_leakage | Mulhu _ _ _ => Return Mulhu_leakage - | Div _ _ _ => Return Div_leakage - | Divu _ _ _ => Return Divu_leakage - | Rem _ _ _ => Return Rem_leakage - | Remu _ _ _ => Return Remu_leakage + | Div _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Div_leakage num den) + | Divu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divu_leakage num den) + | Rem _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Rem_leakage num den) + | Remu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remu_leakage num den) | InvalidM => Return InvalidM_leakage end. @@ -141,15 +144,15 @@ Inductive LeakageI64 | Ld_leakage (addr: word) | Lwu_leakage (addr: word) | Addiw_leakage -| Slliw_leakage -| Srliw_leakage -| Sraiw_leakage +| Slliw_leakage (shamt : Z) +| Srliw_leakage (shamt : Z) +| Sraiw_leakage (shamt : Z) | Sd_leakage (addr: word) | Addw_leakage | Subw_leakage -| Sllw_leakage -| Srlw_leakage -| Sraw_leakage +| Sllw_leakage (shamt : word) +| Srlw_leakage (shamt : word) +| Sraw_leakage (shamt : word) | InvalidI64_leakage. Definition leakage_of_instr_I64 @@ -159,15 +162,15 @@ Definition leakage_of_instr_I64 | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (word.add rs1_val (word.of_Z oimm12)))) | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (word.add rs1_val (word.of_Z oimm12)))) | Addiw _ _ _ => Return Addiw_leakage - | Slliw _ _ _ => Return Slliw_leakage - | Srliw _ _ _ => Return Srliw_leakage - | Sraiw _ _ _ => Return Sraiw_leakage + | Slliw _ _ shamt => Return (Slliw_leakage shamt) + | Srliw _ _ shamt => Return (Srliw_leakage shamt) + | Sraiw _ _ shamt => Return (Sraiw_leakage shamt) | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (word.add rs1_val (word.of_Z simm12)))) | Addw _ _ _ => Return Addw_leakage | Subw _ _ _ => Return Subw_leakage - | Sllw _ _ _ => Return Sllw_leakage - | Srlw _ _ _ => Return Srlw_leakage - | Sraw _ _ _ => Return Sraw_leakage + | Sllw _ _ rs2 => shamt <- getRegister rs2; Return (Sllw_leakage shamt) + | Srlw _ _ rs2 => shamt <- getRegister rs2; Return (Srlw_leakage shamt) + | Sraw _ _ rs2 => shamt <- getRegister rs2; Return (Sraw_leakage shamt) | InvalidI64 => Return InvalidI64_leakage end. @@ -182,26 +185,26 @@ Inductive LeakageI | Fence_leakage (* unsure about this one. *) | Fence_i_leakage | Addi_leakage -| Slli_leakage -| Slti_leakage -| Sltiu_leakage +| Slli_leakage (shamt : Z) +| Slti_leakage (shamt : Z) +| Sltiu_leakage (shamt : Z) | Xori_leakage | Ori_leakage | Andi_leakage -| Srli_leakage -| Srai_leakage +| Srli_leakage (shamt : Z) +| Srai_leakage (shamt : Z) | Auipc_leakage | Sb_leakage (addr: word) | Sh_leakage (addr: word) | Sw_leakage (addr: word) | Add_leakage | Sub_leakage -| Sll_leakage -| Slt_leakage -| Sltu_leakage +| Sll_leakage (shamt : word) +| Slt_leakage (shamt : word) +| Sltu_leakage (shamt : word) | Xor_leakage -| Srl_leakage -| Sra_leakage +| Srl_leakage (shamt : word) +| Sra_leakage (shamt : word) | Or_leakage | And_leakage | Lui_leakage @@ -277,9 +280,9 @@ Inductive InstructionI : Type := | Jalr (rd : Register) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) : InstructionI | Jal (rd : Register) (jimm20 : Utility.Utility.MachineInt) : InstructionI - | InvalidI : InstructionI. Search Z. Compute (Utility.regToZ_unsigned 0). + | InvalidI : InstructionI. -(* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) Print word.word. +(* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) Definition leakage_of_instr_I {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) (instr : InstructionI) : M LeakageI := @@ -292,26 +295,26 @@ Definition leakage_of_instr_I | Fence _ _ => Return Fence_leakage | Fence_i => Return Fence_i_leakage | Addi _ _ _ => Return Addi_leakage - | Slli _ _ _ => Return Slli_leakage - | Slti _ _ _ => Return Slti_leakage - | Sltiu _ _ _ => Return Sltiu_leakage + | Slli _ _ shamt => Return (Slli_leakage shamt) + | Slti _ _ shamt => Return (Slti_leakage shamt) + | Sltiu _ _ shamt => Return (Sltiu_leakage shamt) | Xori _ _ _ => Return Xori_leakage | Ori _ _ _ => Return Ori_leakage | Andi _ _ _ => Return Andi_leakage - | Srli _ _ _ => Return Srli_leakage - | Srai _ _ _ => Return Srai_leakage + | Srli _ _ shamt => Return (Srli_leakage shamt) + | Srai _ _ shamt => Return (Srai_leakage shamt) | Auipc _ _ => Return Auipc_leakage | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (word.add rs1_val (word.of_Z simm12))) | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (word.add rs1_val (word.of_Z simm12))) | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (word.add rs1_val (word.of_Z simm12))) | Add _ _ _ => Return Add_leakage | Sub _ _ _ => Return Sub_leakage - | Sll _ _ _ => Return Sll_leakage - | Slt _ _ _ => Return Slt_leakage - | Sltu _ _ _ => Return Sltu_leakage + | Sll _ _ rs2 => shamt <- getRegister rs2; Return (Sll_leakage shamt) + | Slt _ _ rs2 => shamt <- getRegister rs2; Return (Slt_leakage shamt) + | Sltu _ _ rs2 => shamt <- getRegister rs2; Return (Sltu_leakage shamt) | Xor _ _ _ => Return Xor_leakage - | Srl _ _ _ => Return Srl_leakage - | Sra _ _ _ => Return Sra_leakage + | Srl _ _ rs2 => shamt <- getRegister rs2; Return (Srl_leakage shamt) + | Sra _ _ rs2 => shamt <- getRegister rs2; Return (Sra_leakage shamt) | Or _ _ _ => Return Or_leakage | And _ _ _ => Return And_leakage | Lui _ _ => Return Lui_leakage @@ -693,11 +696,11 @@ Definition leakage_of_instr (instr : Instruction) : M LeakageEvent := match instr with | IInstruction instr => l <- leakage_of_instr_I getRegister instr; Return (ILeakage l) - | MInstruction instr => l <- leakage_of_instr_M instr; Return (MLeakage l) + | MInstruction instr => l <- leakage_of_instr_M getRegister instr; Return (MLeakage l) | AInstruction instr => l <- leakage_of_instr_A getRegister instr; Return (ALeakage l) | FInstruction instr => l <- leakage_of_instr_F getRegister instr; Return (FLeakage l) | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; Return (I64Leakage l) - | M64Instruction instr => l <- leakage_of_instr_M64 instr; Return (M64Leakage l) + | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; Return (M64Leakage l) | A64Instruction instr => l <- leakage_of_instr_A64 getRegister instr; Return (A64Leakage l) | F64Instruction instr => l <- leakage_of_instr_F64 instr; Return (F64Leakage l) | CSRInstruction instr => l <- leakage_of_instr_CSR instr; Return (CSRLeakage l) From a0eecd59298b8f7cb1e363645c85742871437f6c Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Tue, 12 Dec 2023 20:15:10 -0500 Subject: [PATCH 20/42] oops, turns out slt stands for "store less than", nothing to do with "shift logical" --- src/riscv/Spec/Decode.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index de3028e..25ff8fd 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -186,8 +186,8 @@ Inductive LeakageI | Fence_i_leakage | Addi_leakage | Slli_leakage (shamt : Z) -| Slti_leakage (shamt : Z) -| Sltiu_leakage (shamt : Z) +| Slti_leakage +| Sltiu_leakage | Xori_leakage | Ori_leakage | Andi_leakage @@ -200,8 +200,8 @@ Inductive LeakageI | Add_leakage | Sub_leakage | Sll_leakage (shamt : word) -| Slt_leakage (shamt : word) -| Sltu_leakage (shamt : word) +| Slt_leakage +| Sltu_leakage | Xor_leakage | Srl_leakage (shamt : word) | Sra_leakage (shamt : word) @@ -296,8 +296,8 @@ Definition leakage_of_instr_I | Fence_i => Return Fence_i_leakage | Addi _ _ _ => Return Addi_leakage | Slli _ _ shamt => Return (Slli_leakage shamt) - | Slti _ _ shamt => Return (Slti_leakage shamt) - | Sltiu _ _ shamt => Return (Sltiu_leakage shamt) + | Slti _ _ _ => Return Slti_leakage + | Sltiu _ _ _ => Return Sltiu_leakage | Xori _ _ _ => Return Xori_leakage | Ori _ _ _ => Return Ori_leakage | Andi _ _ _ => Return Andi_leakage @@ -310,8 +310,8 @@ Definition leakage_of_instr_I | Add _ _ _ => Return Add_leakage | Sub _ _ _ => Return Sub_leakage | Sll _ _ rs2 => shamt <- getRegister rs2; Return (Sll_leakage shamt) - | Slt _ _ rs2 => shamt <- getRegister rs2; Return (Slt_leakage shamt) - | Sltu _ _ rs2 => shamt <- getRegister rs2; Return (Sltu_leakage shamt) + | Slt _ _ _ => Return Slt_leakage + | Sltu _ _ _ => Return Sltu_leakage | Xor _ _ _ => Return Xor_leakage | Srl _ _ rs2 => shamt <- getRegister rs2; Return (Srl_leakage shamt) | Sra _ _ rs2 => shamt <- getRegister rs2; Return (Sra_leakage shamt) From c161084ad928e5e2eb55b566512d935b57f56814 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Thu, 28 Dec 2023 11:43:32 -0500 Subject: [PATCH 21/42] not sure --- src/riscv/Spec/Machine.v | 1 + 1 file changed, 1 insertion(+) diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index db4d734..84ec309 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -29,6 +29,7 @@ Definition PrivMode_eqb(m1 m2: PrivMode): bool := Z.eqb (encodePrivMode m1) (enc Definition PrivMode_ltb(m1 m2: PrivMode): bool := Z.ltb (encodePrivMode m1) (encodePrivMode m2). Inductive AccessType: Set := Instr | Load | Store. +(*find usages of Instr for logic about instruction fetching*) Inductive SourceType: Set := VirtualMemory | Fetch | Execute. Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { From d840c1a452bf2551940fb0f07ae99e82ce43a3e2 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Fri, 12 Jul 2024 05:10:39 -0400 Subject: [PATCH 22/42] anonymize --- src/riscv/Spec/Decode.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 25ff8fd..2a28dcc 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -62,7 +62,7 @@ Inductive InstructionM64 : Type := Inductive LeakageM64 {width} {BW : Bitwidth width} {word: word.word width} : Type := -| Mulw_leakage (*andres says that this is constant-time*) +| Mulw_leakage | Divw_leakage (num : word) (den : word) (*but that this is not.*) | Divuw_leakage (num : word) (den : word) | Remw_leakage (num : word) (den : word) (*i'm not sure about this one. probably same as div, so not constant-time.*) @@ -208,7 +208,7 @@ Inductive LeakageI | Or_leakage | And_leakage | Lui_leakage -| Beq_leakage (branch: bool) (* unsure whether this should be here - i think that andres said that having this argument would make my life harder *) +| Beq_leakage (branch: bool) | Bne_leakage (branch: bool) | Blt_leakage (branch: bool) | Bge_leakage (branch: bool) From 3a47d9decd924b2c238ecd97f6b1f0277a0fc0a4 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Fri, 2 Aug 2024 14:18:14 -0400 Subject: [PATCH 23/42] clean things up: delete print statements, fix whitespace, fix merging errors, etc. --- src/riscv/Platform/MaterializeRiscvProgram.v | 10 +- src/riscv/Platform/MetricMinimal.v | 6 +- src/riscv/Platform/MetricMinimalMMIO.v | 36 +- src/riscv/Platform/MetricMinimalNoMul.v | 147 -------- src/riscv/Platform/Minimal.v | 16 +- src/riscv/Platform/MinimalMMIO_Post.v | 6 +- src/riscv/Platform/MinimalNoMul.v | 370 ------------------- src/riscv/Platform/Run.v | 4 +- src/riscv/Spec/Machine.v | 5 +- src/riscv/Spec/MetricPrimitives.v | 11 +- src/riscv/Spec/Primitives.v | 4 + 11 files changed, 27 insertions(+), 588 deletions(-) delete mode 100644 src/riscv/Platform/MetricMinimalNoMul.v delete mode 100644 src/riscv/Platform/MinimalNoMul.v diff --git a/src/riscv/Platform/MaterializeRiscvProgram.v b/src/riscv/Platform/MaterializeRiscvProgram.v index 36de0e1..d1065ae 100644 --- a/src/riscv/Platform/MaterializeRiscvProgram.v +++ b/src/riscv/Platform/MaterializeRiscvProgram.v @@ -84,12 +84,12 @@ Section Riscv. setPC a := act (SetPC a) ret; endCycleNormal := act EndCycleNormal ret; endCycleEarly A := act (EndCycleEarly A) ret; - |}. + |}. - Global Instance MaterializeWithLeakage : RiscvProgramWithLeakage := - {| - RVP := Materialize; - leakEvent a := act (LeakEvent a) ret |}. + Global Instance MaterializeWithLeakage : RiscvProgramWithLeakage := {| + RVP := Materialize; + leakEvent a := act (LeakEvent a) ret + |}. (* Not (yet) in Riscv monad, but added here because it's useful to initialize nextPc to pc+4 at the beginning of each cycle instead of having to maintain diff --git a/src/riscv/Platform/MetricMinimal.v b/src/riscv/Platform/MetricMinimal.v index e8fcd9a..f6e1c60 100644 --- a/src/riscv/Platform/MetricMinimal.v +++ b/src/riscv/Platform/MetricMinimal.v @@ -64,12 +64,12 @@ Section Riscv. fence := liftL2 id fence; endCycleNormal := liftL0 (addMetricInstructions 1) endCycleNormal; endCycleEarly{A} := liftL0 (addMetricInstructions 1) (@endCycleEarly _ _ _ _ _ A); - }. + }. - Instance IsMetricRiscvMachineWithLeakage: @RiscvProgramWithLeakage width BW word (OState MetricRiscvMachine) _ _ _ := { + Instance IsMetricRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (OState MetricRiscvMachine) _ _ _ := { RVP := IsMetricRiscvMachine; leakEvent := liftL1 id leakEvent; - }. + }. Arguments Memory.load_bytes: simpl never. Arguments Memory.store_bytes: simpl never. diff --git a/src/riscv/Platform/MetricMinimalMMIO.v b/src/riscv/Platform/MetricMinimalMMIO.v index 1090c2a..ac0318a 100644 --- a/src/riscv/Platform/MetricMinimalMMIO.v +++ b/src/riscv/Platform/MetricMinimalMMIO.v @@ -30,41 +30,7 @@ Section Riscv. (* note: ext_spec does not have access to the metrics *) Context {mmio_spec: MMIOSpec}. - (*should i change this to (list LeakageEvent -> list LeakageEvent) * ... ? *) - Definition action : Type := (MetricLog -> MetricLog) * riscv_primitive. - Definition result (a : action) := primitive_result (snd a). - Local Notation M := (free action result). Print act. - - Global Instance IsRiscvMachine: RiscvProgram M word := {| - getRegister a := act (id, GetRegister a) ret; - setRegister a b := act (id, SetRegister a b) ret; - loadByte a b := act (addMetricLoads 1, LoadByte a b) ret; - loadHalf a b := act (addMetricLoads 1, LoadHalf a b) ret; - loadWord a b := act (addMetricLoads 1, LoadWord a b) ret; - loadDouble a b := act (addMetricLoads 1, LoadDouble a b) ret; - storeByte a b c := act (addMetricStores 1, StoreByte a b c) ret; - storeHalf a b c := act (addMetricStores 1, StoreHalf a b c) ret; - storeWord a b c := act (addMetricStores 1, StoreWord a b c) ret; - storeDouble a b c := act (addMetricStores 1, StoreDouble a b c) ret; - makeReservation a := act (id, MakeReservation a) ret; - clearReservation a := act (id, ClearReservation a) ret; - checkReservation a := act (id, CheckReservation a) ret; - getCSRField f := act (id, GetCSRField f) ret; - setCSRField f v := act (id, SetCSRField f v) ret; - getPrivMode := act (id, GetPrivMode) ret; - setPrivMode m := act (id, SetPrivMode m) ret; - fence a b := act (id, Fence a b) ret; - getPC := act (id, GetPC) ret; - setPC a := act (addMetricJumps 1, SetPC a) ret; - endCycleNormal := act (addMetricInstructions 1, EndCycleNormal) ret; - endCycleEarly A := act (addMetricInstructions 1, EndCycleEarly A) ret; - |}. - - Check addMetricInstructions. - - Global Instance IsRiscvMachineWithLeakage: RiscvProgramWithLeakage := - {| - leakEvent a := act (id, LeakEvent a) ret; |}. + Local Notation M := (free action result). Definition interp_action a metmach post := interpret_action (snd a) (metmach.(getMachine)) (fun r mach => diff --git a/src/riscv/Platform/MetricMinimalNoMul.v b/src/riscv/Platform/MetricMinimalNoMul.v deleted file mode 100644 index 1c2acfb..0000000 --- a/src/riscv/Platform/MetricMinimalNoMul.v +++ /dev/null @@ -1,147 +0,0 @@ -Require Import Coq.ZArith.ZArith. -Require Import Coq.Logic.FunctionalExtensionality. -Require Import Coq.Logic.PropExtensionality. -Require Import riscv.Utility.Monads. -Require Import riscv.Utility.MonadNotations. -Require Import riscv.Spec.Decode. -Require Import riscv.Spec.Machine. -Require Import riscv.Utility.Utility. -Require Import riscv.Spec.Primitives. -Require Import riscv.Spec.MetricPrimitives. -Require Import Coq.Lists.List. Import ListNotations. -Require Export riscv.Platform.RiscvMachine. -Require Export riscv.Platform.MetricRiscvMachine. -Require Import riscv.Platform.MinimalNoMul. -Require Import riscv.Platform.MetricLogging. -Require Import coqutil.Z.Lia. -Require Import coqutil.Map.Interface. -Require Import coqutil.Tactics.Tactics. -Import MetricRiscvMachine. - -Local Open Scope Z_scope. -Local Open Scope bool_scope. - -Section Riscv. - Import List. - Import free. - - Context {width: Z} {BW: Bitwidth width} {word: word width} {word_ok: word.ok word}. - Context {Mem: map.map word byte}. - Context {Registers: map.map Register word}. - - Local Notation M := (free riscv_primitive primitive_result). - - Definition metrics_of(p: riscv_primitive): MetricLog -> MetricLog := - match p with - | GetRegister a => id - | SetRegister a b => id - | LoadByte a b => addMetricLoads 1 - | LoadHalf a b => addMetricLoads 1 - | LoadWord a b => addMetricLoads 1 - | LoadDouble a b => addMetricLoads 1 - | StoreByte a b c => addMetricStores 1 - | StoreHalf a b c => addMetricStores 1 - | StoreWord a b c => addMetricStores 1 - | StoreDouble a b c => addMetricStores 1 - | MakeReservation a => id - | ClearReservation a => id - | CheckReservation a => id - | GetCSRField f => id - | SetCSRField f v => id - | GetPrivMode => id - | SetPrivMode m => id - | Fence a b => id - | LeakEvent a => id - | GetPC => id - | SetPC a => addMetricJumps 1 - | StartCycle => id - | EndCycleNormal => addMetricInstructions 1 - | EndCycleEarly A => addMetricInstructions 1 - end. - - Definition interp_action p (m: MetricRiscvMachine) post := - interpret_action p m.(getMachine) - (fun r mach => post r (mkMetricRiscvMachine mach (metrics_of p m.(getMetrics)))) - (fun _ => False). - - Arguments Memory.load_bytes: simpl never. - Arguments Memory.store_bytes: simpl never. - Arguments LittleEndian.combine: simpl never. - - Global Instance MetricMinimalNoMulPrimitivesParams: PrimitivesParams M MetricRiscvMachine := - { - Primitives.mcomp_sat := @free.interp _ _ _ interp_action; - Primitives.is_initial_register_value x := True; - Primitives.nonmem_load := Primitives.nonmem_load (PrimitivesParams := MinimalNoMulPrimitivesParams); - Primitives.nonmem_store := Primitives.nonmem_store (PrimitivesParams := MinimalNoMulPrimitivesParams); - Primitives.valid_machine mach := no_M mach.(getMachine); - }. - - Global Instance MinimalNoMulSatisfies_mcomp_sat_spec: mcomp_sat_spec MetricMinimalNoMulPrimitivesParams. - Proof. - split; cbv [mcomp_sat MetricMinimalNoMulPrimitivesParams Monad_free Bind Return]. - { symmetry. eapply interp_bind_ex_mid; intros. - eapply MinimalNoMul.interpret_action_weaken_post; eauto; cbn; eauto. } - { symmetry. rewrite interp_ret; eapply iff_refl. } - Qed. - - Lemma interp_action_weaken_post a (post1 post2:_->_->Prop) - (H: forall r s, post1 r s -> post2 r s) s - : interp_action a s post1 -> interp_action a s post2. - Proof. eapply MinimalNoMul.interpret_action_weaken_post; eauto. Qed. - Lemma interp_action_appendonly' a s post : - interp_action a s post -> - interp_action a s (fun v s' => post v s' /\ endswith s'.(getLog) s.(getLog)). - Proof. eapply MinimalNoMul.interpret_action_appendonly''; eauto. Qed. - Lemma interp_action_total{memOk: map.ok Mem} a (s: MetricRiscvMachine) post : - no_M s -> - interp_action a s post -> - exists v s, post v s /\ no_M s. - Proof. - intros H H1. - unshelve epose proof (MinimalNoMul.interpret_action_total _ _ _ _ _ H1) as H0; eauto. - destruct H0 as (?&?&[[]|(?&?)]); eauto. - Qed. - Lemma interp_action_preserves_valid{memOk: map.ok Mem} a s post : - no_M s.(getMachine) -> - interp_action a s post -> - interp_action a s (fun v s' => post v s' /\ no_M s'.(getMachine)). - Proof. - intros D I. - unshelve epose proof (MinimalNoMul.interpret_action_preserves_valid' _ _ _ D I) as H0; eauto. - Qed. Search RiscvProgramWithLeakage. Print MetricPrimitivesSane. Search RiscvProgramWithLeakage. - - Global Instance MetricMinimalNoMulPrimitivesSane{memOk: map.ok Mem} : - MetricPrimitivesSane MetricMinimalNoMulPrimitivesParams. - Proof. - split; cbv [mcomp_sane valid_machine MetricMinimalNoMulPrimitivesParams]; - intros *; intros D M; - (split; [ exact (interp_action_total _ st _ D M) - | eapply interp_action_preserves_valid; try eassumption; - eapply interp_action_appendonly'; try eassumption ]). - Qed. - - Global Instance MetricMinimalNoMulSatisfiesPrimitives{memOk: map.ok Mem}: - MetricPrimitives MetricMinimalNoMulPrimitivesParams. - Proof. - split; try exact _. - all : cbv [mcomp_sat spec_load spec_store MetricMinimalNoMulPrimitivesParams invalidateWrittenXAddrs]. - all: intros; destruct initialL; - repeat match goal with - | _ => progress subst - | _ => Option.inversion_option - | _ => progress cbn -[Memory.load_bytes Memory.store_bytes HList.tuple] in * - | _ => progress cbv [id valid_register is_initial_register_value load store Memory.loadByte Memory.loadHalf Memory.loadWord Memory.loadDouble Memory.storeByte Memory.storeHalf Memory.storeWord Memory.storeDouble] in * - | H : exists _, _ |- _ => destruct H - | H : _ /\ _ |- _ => destruct H - | |- _ => solve [ intuition (eauto || blia) ] - | H : _ \/ _ |- _ => destruct H - | |- context[match ?x with _ => _ end] => destruct x eqn:? - | |- _ => progress unfold getReg, setReg - | |-_ /\ _ => split - end. - (* setRegister *) - 1: destruct getMachine; eassumption. - Qed. - -End Riscv. diff --git a/src/riscv/Platform/Minimal.v b/src/riscv/Platform/Minimal.v index c6f39b7..cdfbb44 100644 --- a/src/riscv/Platform/Minimal.v +++ b/src/riscv/Platform/Minimal.v @@ -100,11 +100,10 @@ Section Riscv. endCycleEarly{A: Type} := fail_hard; }. - Print RiscvProgramWithLeakage. - Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage width BW word (OState RiscvMachine) _ _ _ := { + Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (OState RiscvMachine) _ _ _ := { RVP := IsRiscvMachine; leakEvent e := fail_hard; - }. + }. Arguments Memory.load_bytes: simpl never. Arguments Memory.store_bytes: simpl never. @@ -231,21 +230,12 @@ Section Riscv. intros. eapply update_sane. intros. exists [e]. destruct mach. reflexivity. Qed. - Print PrimitivesSane. - Instance MinimalSane: PrimitivesSane MinimalPrimitivesParams. Proof. constructor. all: intros; unfold IsRiscvMachine, IsRiscvMachineWithLeakage, RVP, - getRegister, setRegister, - loadByte, loadHalf, loadWord, loadDouble, - storeByte, storeHalf, storeWord, storeDouble, - getPC, setPC, - endCycleNormal, endCycleEarly, raiseExceptionWithInfo, - loadN, storeN, fail_if_None. - all: - unfold getRegister, setRegister, + getRegister, setRegister, loadByte, loadHalf, loadWord, loadDouble, storeByte, storeHalf, storeWord, storeDouble, getPC, setPC, diff --git a/src/riscv/Platform/MinimalMMIO_Post.v b/src/riscv/Platform/MinimalMMIO_Post.v index 9d1a8f3..f362135 100644 --- a/src/riscv/Platform/MinimalMMIO_Post.v +++ b/src/riscv/Platform/MinimalMMIO_Post.v @@ -115,10 +115,10 @@ Section Riscv. - exact (post tt (withPc mach.(getNextPc) (withNextPc (word.add mach.(getNextPc) (word.of_Z 4)) mach))). Defined. - Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage width BW word (Post RiscvMachine) _ _ _ := - {| + Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (Post RiscvMachine) _ _ _ := {| RVP := IsRiscvMachine; - leakEvent _ := fun _ _ => False |}. + leakEvent _ := fun _ _ => False; + |}. Definition MinimalMMIOPrimitivesParams: PrimitivesParams (Post RiscvMachine) RiscvMachine := {| Primitives.mcomp_sat A (m: Post RiscvMachine A) mach post := m mach post; diff --git a/src/riscv/Platform/MinimalNoMul.v b/src/riscv/Platform/MinimalNoMul.v deleted file mode 100644 index f588c0d..0000000 --- a/src/riscv/Platform/MinimalNoMul.v +++ /dev/null @@ -1,370 +0,0 @@ -Require Import Coq.Strings.String. -Require Import Coq.ZArith.ZArith. -Require Import riscv.Utility.Monads. -Require Import riscv.Utility.MonadNotations. -Require Export riscv.Utility.FreeMonad. -Require Import riscv.Spec.Decode. -Require Import riscv.Spec.Machine. -Require Import riscv.Utility.Utility. -Require Import riscv.Spec.Primitives. -Require Import Coq.Lists.List. Import ListNotations. -Require Import coqutil.Datatypes.List. -Require Import coqutil.Datatypes.ListSet. -Require Export riscv.Platform.RiscvMachine. -Require Export riscv.Platform.MaterializeRiscvProgram. -Require Import coqutil.Z.Lia. -Require Import coqutil.Map.Interface. -Require Import coqutil.Map.Properties. -Require Import coqutil.Word.Properties. -Require Import coqutil.Datatypes.PropSet. -Require Import coqutil.Tactics.Tactics. -Require Import coqutil.Tactics.fwd. -Require Import riscv.Platform.Sane. - -Local Open Scope Z_scope. -Local Open Scope bool_scope. - -Section Riscv. - Import free. - Context {width: Z} {BW: Bitwidth width} {word: word width} {word_ok: word.ok word}. - Context {Mem: map.map word byte} {Registers: map.map Register word}. - - Add Ring wring : (word.ring_theory (word := word)) - (preprocess [autorewrite with rew_word_morphism], - morphism (word.ring_morph (word := word)), - constants [word_cst]). - - Definition store(n: nat)(ctxid: SourceType) a v mach post := - match Memory.store_bytes n mach.(getMem) a v with - | Some m => post (withXAddrs (invalidateWrittenXAddrs n a mach.(getXAddrs)) (withMem m mach)) - | None => False - end. - - Definition load(n: nat)(ctxid: SourceType) a mach post := - (ctxid = Fetch -> isXAddr4 a mach.(getXAddrs)) /\ - match Memory.load_bytes n mach.(getMem) a with - | Some v => post v mach - | None => False - end. - - Definition updatePc(mach: RiscvMachine): RiscvMachine := - withPc mach.(getNextPc) (withNextPc (word.add mach.(getNextPc) (word.of_Z 4)) mach). - - Definition getReg(regs: Registers)(reg: Z): word := - if ((0 x - | None => word.of_Z 0 - end - else word.of_Z 0. - - Definition setReg(reg: Z)(v: word)(regs: Registers): Registers := - if ((0 RiscvMachine -> Prop) -> (RiscvMachine -> Prop) -> Prop := - match a with - | GetRegister reg => fun postF postA => - let v := getReg mach.(getRegs) reg in - postF v mach - | SetRegister reg v => fun postF postA => - let regs := setReg reg v mach.(getRegs) in - postF tt (withRegs regs mach) - | GetPC => fun postF postA => postF mach.(getPc) mach - | SetPC newPC => fun postF postA => postF tt (withNextPc newPC mach) - | LoadByte ctxid a => fun postF postA => load 1 ctxid a mach postF - | LoadHalf ctxid a => fun postF postA => load 2 ctxid a mach postF - | LoadWord ctxid a => fun postF postA => load 4 ctxid a mach postF - | LoadDouble ctxid a => fun postF postA => load 8 ctxid a mach postF - | StoreByte ctxid a v => fun postF postA => store 1 ctxid a v mach (postF tt) - | StoreHalf ctxid a v => fun postF postA => store 2 ctxid a v mach (postF tt) - | StoreWord ctxid a v => fun postF postA => store 4 ctxid a v mach (postF tt) - | StoreDouble ctxid a v => fun postF postA => store 8 ctxid a v mach (postF tt) - | StartCycle => fun postF postA => - postF tt (withNextPc (word.add mach.(getPc) (word.of_Z 4)) mach) - | EndCycleNormal => fun postF postA => postF tt (updatePc mach) - | EndCycleEarly _ => fun postF postA => postA (updatePc mach) (* ignores postF containing the continuation *) - | LeakEvent e => fun postF postA => postF tt (withLeakageEvent e mach) - | MakeReservation _ - | ClearReservation _ - | CheckReservation _ - | GetCSRField _ - | SetCSRField _ _ - | GetPrivMode - | SetPrivMode _ - | Fence _ _ - => fun postF postA => False - end. - - Definition no_M(mach: RiscvMachine): Prop := - forall a v, - isXAddr4 a mach.(getXAddrs) -> - word.unsigned a mod 4 = 0 -> - Memory.load_bytes 4 mach.(getMem) a = Some v -> - forall minst, decode RV32IM (LittleEndian.combine 4 v) <> MInstruction minst. - - Instance MinimalNoMulPrimitivesParams: - PrimitivesParams (free riscv_primitive primitive_result) RiscvMachine := - {| - Primitives.mcomp_sat A m mach postF := - @free.interpret _ _ _ interpret_action A m mach postF (fun _ => False); - Primitives.is_initial_register_value x := True; - Primitives.nonmem_load _ _ _ _ _ := False; - Primitives.nonmem_store _ _ _ _ _ _ := False; - Primitives.valid_machine := no_M; - |}. - - Lemma load_weaken_post n c a m (post1 post2:_->_->Prop) - (H: forall r s, post1 r s -> post2 r s) - : load n c a m post1 -> load n c a m post2. - Proof. - cbv [load nonmem_load]. - destruct (Memory.load_bytes n (getMem m) a); intuition eauto. - Qed. - - Lemma store_weaken_post n c a v m (post1 post2:_->Prop) - (H: forall s, post1 s -> post2 s) - : store n c a v m post1 -> store n c a v m post2. - Proof. - cbv [store nonmem_store]. - destruct (Memory.store_bytes n (getMem m) a); intuition eauto. - Qed. - - Lemma interpret_action_weaken_post a (postF1 postF2: _ -> _ -> Prop) (postA1 postA2: _ -> Prop): - (forall r s, postF1 r s -> postF2 r s) -> - (forall s, postA1 s -> postA2 s) -> - forall s, interpret_action a s postF1 postA1 -> interpret_action a s postF2 postA2. - Proof. - destruct a; cbn; try solve [intuition eauto]. - all : eauto using load_weaken_post, store_weaken_post. - Qed. - - Global Instance MinimalNoMulSatisfies_mcomp_sat_spec: mcomp_sat_spec MinimalNoMulPrimitivesParams. - Proof. - split; cbv [mcomp_sat MinimalNoMulPrimitivesParams Bind Return Monad_free]. - { symmetry. eapply interpret_bind_ex_mid, interpret_action_weaken_post. } - { symmetry; intros. rewrite interpret_ret; eapply iff_refl. } - Qed. - - Lemma preserve_undef_on{memOk: map.ok Mem}: forall n (m m': Mem) a w s, - Memory.store_bytes n m a w = Some m' -> - map.undef_on m s -> - map.undef_on m' s. - Proof. - eauto using map.same_domain_preserves_undef_on, Memory.store_bytes_preserves_domain. - Qed. - - Lemma removeXAddr_bw: forall (a1 a2: word) xaddrs, - isXAddr1 a1 (removeXAddr a2 xaddrs) -> - isXAddr1 a1 xaddrs. - Proof. - unfold isXAddr1, removeXAddr. - intros. - eapply filter_In. - eassumption. - Qed. - - Lemma invalidateWrittenXAddrs_bw: forall n (a r: word) xa, - In a (invalidateWrittenXAddrs n r xa) -> - In a xa. - Proof. - induction n; cbn; intros. - - assumption. - - eapply IHn. eapply removeXAddr_bw. unfold isXAddr1. eassumption. - Qed. - - Lemma put_preserves_getmany_of_tuple{memOk: map.ok Mem}: - forall n (t: HList.tuple word n) (m: Mem) (r: word) b, - ~In r (HList.tuple.to_list t) -> - map.getmany_of_tuple m t = - map.getmany_of_tuple (map.put m r b) t. - Proof. - induction n; intros. - - destruct t. reflexivity. - - destruct t as (w & t). cbn in H|-*. - unfold map.getmany_of_tuple in IHn. - erewrite IHn. - 2: { - intro C. eapply H. right. exact C. - } - rewrite ?map.get_put_dec. - destr (word.eqb r w). 2: reflexivity. - exfalso. apply H. auto. - Qed. - - Lemma transfer_load4bytes_to_previous_mem{memOk: map.ok Mem}: - forall n (a: word) v m m' r w someSet, - Memory.store_bytes n m r w = Some m' -> - (* a notin r..r+n *) - isXAddr4 a (invalidateWrittenXAddrs n r someSet) -> - Memory.load_bytes 4 m' a = Some v -> - Memory.load_bytes 4 m a = Some v. - Proof. - induction n; intros. - - cbn in H. congruence. - - unfold Memory.store_bytes in *. fwd. cbn in H0. destruct w as [b w]. - cbn -[HList.tuple Memory.load_bytes] in H1. - cbn in E. fwd. - unfold Memory.load_bytes at 1 in IHn. - unfold map.getmany_of_tuple, Memory.footprint in IHn. - specialize IHn with (m := m) (r := (word.add r (word.of_Z 1))). - rewrite E1 in IHn. - eapply IHn. - + reflexivity. - + instantiate (1 := someSet). clear -H0. - unfold isXAddr4 in *. fwd. eauto 10 using removeXAddr_bw. - + unfold Memory.load_bytes in *. - etransitivity. 2: eassumption. - unfold Memory.unchecked_store_bytes, Memory.footprint. - eapply put_preserves_getmany_of_tuple. - cbn. clear -H0 word_ok. unfold isXAddr4, isXAddr1 in H0. fwd. - intro C. - unfold removeXAddr in *. - apply_in_hyps filter_In. - fwd. - apply_in_hyps Bool.negb_true_iff. - apply_in_hyps Properties.word.eqb_false. - repeat destruct C as [C | C]; try assumption; - match type of C with - | ?l = _ => ring_simplify l in C - end; - subst r; - congruence. - Qed. - - Lemma isXAddr4_uninvalidate: forall (a: word) n r xa, - isXAddr4 a (invalidateWrittenXAddrs n r xa) -> - isXAddr4 a xa. - Proof. - unfold isXAddr4, isXAddr1. intros. fwd. eauto 10 using invalidateWrittenXAddrs_bw. - Qed. - - Lemma interpret_action_total{memOk: map.ok Mem} a s postF postA : - no_M s -> - interpret_action a s postF postA -> - exists s', no_M s' /\ (postA s' \/ exists v', postF v' s'). - Proof. - destruct s, a; cbn -[HList.tuple Memory.load_bytes invalidateWrittenXAddrs]; - cbv [load store no_M]; cbn -[HList.tuple Memory.load_bytes invalidateWrittenXAddrs]; - repeat destruct_one_match; - intuition idtac; - repeat lazymatch goal with - | H : postF _ ?mach |- exists _ : RiscvMachine, _ => - exists mach; cbn [RiscvMachine.getMem RiscvMachine.getXAddrs] - | H : postA ?mach |- exists _ : RiscvMachine, _ => - exists mach; cbn [RiscvMachine.getMem RiscvMachine.getXAddrs] - | Hexists : (exists v, ?P), Hforall : (forall v, ?P -> _) |- _ => - let v := fresh "v" in - destruct Hexists as [v Hexists]; - specialize (Hforall v Hexists) - end; - ssplit; eauto; cbn -[HList.tuple Memory.load_bytes invalidateWrittenXAddrs]; - change removeXAddr with (@List.removeb word word.eqb); - rewrite ?ListSet.of_list_removeb; - intuition eauto 10 using transfer_load4bytes_to_previous_mem, isXAddr4_uninvalidate. - Qed. - - Lemma interpret_action_total'{memOk: map.ok Mem} a s post : - no_M s -> - interpret_action a s post (fun _ : RiscvMachine => False) -> - exists v s', post v s' /\ no_M s'. - Proof. - intros. pose proof interpret_action_total as P. - specialize P with (postA := (fun _ : RiscvMachine => False)). simpl in P. - specialize (P _ _ _ H H0). - destruct P as (s' & ? & ?). - destruct H2 as [[] | (v' & ?)]. - eauto. - Qed. - - Import coqutil.Tactics.Tactics. - - Lemma interpret_action_appendonly a s postF postA : - interpret_action a s postF postA -> - interpret_action a s (fun _ s' => endswith s'.(getLog) s.(getLog)) - (fun s' => endswith s'.(getLog) s.(getLog)). - Proof. - destruct s, a; cbn; cbv [load store nonmem_load nonmem_store]; cbn; - repeat destruct_one_match; - intuition eauto using endswith_refl, endswith_cons_l. - Qed. - - (* NOTE: maybe instead a generic lemma to push /\ into postcondition? *) - Lemma interpret_action_appendonly' a s postF postA : - interpret_action a s postF postA -> - interpret_action a s (fun v s' => postF v s' /\ endswith s'.(getLog) s.(getLog)) - (fun s' => postA s' /\ endswith s'.(getLog) s.(getLog)). - Proof. - destruct s, a; cbn; cbv [load store nonmem_load nonmem_store]; cbn; - repeat destruct_one_match; intros; destruct_products; try split; - intuition eauto using endswith_refl, endswith_cons_l. - Qed. - - Lemma interpret_action_appendonly'' a s post : - interpret_action a s post (fun _ : RiscvMachine => False) -> - interpret_action a s (fun v s' => post v s' /\ endswith s'.(getLog) s.(getLog)) - (fun _ : RiscvMachine => False). - Proof. - intros. pose proof interpret_action_appendonly' as P. - specialize (P _ _ _ (fun _ : RiscvMachine => False) H). simpl in P. - eapply interpret_action_weaken_post. 3: exact P. all: simpl; intuition eauto. - Qed. - - Lemma interpret_action_preserves_valid{memOk: map.ok Mem} a s postF postA : - no_M s -> - interpret_action a s postF postA -> - interpret_action a s (fun v s' => postF v s' /\ no_M s') - (fun s' => postA s' /\ no_M s'). - Proof. - destruct s, a; cbn; cbv [load store no_M]; - cbn -[HList.tuple Memory.load_bytes invalidateWrittenXAddrs]; - repeat destruct_one_match; intros; destruct_products; try split; - change removeXAddr with (@List.removeb word word.eqb); - rewrite ?ListSet.of_list_removeb; - intuition eauto 10 using transfer_load4bytes_to_previous_mem, isXAddr4_uninvalidate. - Qed. - - Lemma interpret_action_preserves_valid'{memOk: map.ok Mem} a s post : - no_M s -> - interpret_action a s post (fun _ : RiscvMachine => False) -> - interpret_action a s (fun v s' => post v s' /\ no_M s') - (fun _ : RiscvMachine => False). - Proof. - intros. pose proof interpret_action_preserves_valid as P. - specialize (P _ _ _ (fun _ : RiscvMachine => False) H H0). simpl in P. - eapply interpret_action_weaken_post. 3: exact P. all: simpl; intuition eauto. - Qed. - - Global Instance MinimalNoMulPrimitivesSane{memOk: map.ok Mem} : - PrimitivesSane MinimalNoMulPrimitivesParams. - Proof. - split; cbv [mcomp_sane valid_machine MinimalNoMulPrimitivesParams]; intros *; intros D M; - (split; [ exact (interpret_action_total' _ st _ D M) - | eapply interpret_action_preserves_valid'; try eassumption; - eapply interpret_action_appendonly''; try eassumption ]). - Qed. - - Global Instance MinimalNoMulSatisfiesPrimitives{memOk: map.ok Mem} : - Primitives MinimalNoMulPrimitivesParams. - Proof. - split; try exact _. - all : cbv [mcomp_sat spec_load spec_store MinimalNoMulPrimitivesParams invalidateWrittenXAddrs]. - all: intros; - repeat match goal with - | _ => progress subst - | _ => Option.inversion_option - | _ => progress cbn -[Memory.load_bytes Memory.store_bytes HList.tuple] - | _ => progress cbv [valid_register is_initial_register_value load store Memory.loadByte Memory.loadHalf Memory.loadWord Memory.loadDouble Memory.storeByte Memory.storeHalf Memory.storeWord Memory.storeDouble] in * - | H : exists _, _ |- _ => destruct H - | H : _ /\ _ |- _ => destruct H - | |- _ => solve [ intuition (eauto || blia) ] - | H : _ \/ _ |- _ => destruct H - | |- context[match ?x with _ => _ end] => destruct x eqn:? - | |- _ => progress unfold getReg, setReg - | |-_ /\ _ => split - end. - (* setRegister *) - destruct initialL; eassumption. - Qed. - -End Riscv. diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index d230bb6..0623084 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -16,10 +16,8 @@ Section Riscv. Context {M: Type -> Type}. Context {MM: Monad M}. Context {RVM: RiscvProgramWithLeakage}. - Context {RVS: RiscvMachine M mword}. Print leakage_of_instr. + Context {RVS: RiscvMachine M mword}. - Check (leakage_of_instr getRegister _). - Print leakage_of_instr. Check getRegister. Definition run1(iset: InstructionSet): M unit := pc <- getPC; diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index 84ec309..96ae1b0 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -29,7 +29,6 @@ Definition PrivMode_eqb(m1 m2: PrivMode): bool := Z.eqb (encodePrivMode m1) (enc Definition PrivMode_ltb(m1 m2: PrivMode): bool := Z.ltb (encodePrivMode m1) (encodePrivMode m2). Inductive AccessType: Set := Instr | Load | Store. -(*find usages of Instr for logic about instruction fetching*) Inductive SourceType: Set := VirtualMemory | Fetch | Execute. Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { @@ -65,7 +64,7 @@ Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { Some instances may not support endCycleEarly and fail instead. *) endCycleNormal: M unit; endCycleEarly: forall A, M A; -}. Print LeakageEvent. +}. Class RiscvProgramWithLeakage {width} {BW : Bitwidth width} {word: word.word width} @@ -73,7 +72,7 @@ Class RiscvProgramWithLeakage mkRiscvProgramWithLeakage { RVP :> @RiscvProgram M t MM MWt; leakEvent : @LeakageEvent width BW word -> M unit; - }. + }. Class RiscvMachine`{MP: RiscvProgram} := mkRiscvMachine { (* checks that addr is aligned, and translates the (possibly virtual) addr to a physical diff --git a/src/riscv/Spec/MetricPrimitives.v b/src/riscv/Spec/MetricPrimitives.v index 3e913a4..8deae7f 100644 --- a/src/riscv/Spec/MetricPrimitives.v +++ b/src/riscv/Spec/MetricPrimitives.v @@ -20,8 +20,7 @@ Section MetricPrimitives. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgramWithLeakage}. Print RiscvMachine. Print RiscvProgramWithLeakage. - Set Printing All. Print RVM. + Context {RVM: RiscvProgramWithLeakage}. Context {RVS: @RiscvMachine M word _ _ RVM.(RVP)}. (* monadic computations used for specifying the behavior of RiscvMachines should be "sane" @@ -115,10 +114,6 @@ Section MetricPrimitives. spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM.(RVP))) Memory.loadWord; spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM.(RVP))) Memory.loadDouble; - spec_leakEvent: forall (initialL: MetricRiscvMachine) (post: unit -> MetricRiscvMachine -> Prop) (e : LeakageEvent), - post tt (withLeakageEvent e initialL) -> - mcomp_sat (leakEvent e) initialL post; - spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM.(RVP))) Memory.storeByte; spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM.(RVP))) Memory.storeHalf; spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM.(RVP))) Memory.storeWord; @@ -140,6 +135,10 @@ Section MetricPrimitives. (updateMetrics (addMetricInstructions 1) initialL))) -> mcomp_sat endCycleNormal initialL post; + + spec_leakEvent: forall (initialL: MetricRiscvMachine) (post: unit -> MetricRiscvMachine -> Prop) (e : LeakageEvent), + post tt (withLeakageEvent e initialL) -> + mcomp_sat (leakEvent e) initialL post; }. End MetricPrimitives. diff --git a/src/riscv/Spec/Primitives.v b/src/riscv/Spec/Primitives.v index 09b5a5c..911c7b8 100644 --- a/src/riscv/Spec/Primitives.v +++ b/src/riscv/Spec/Primitives.v @@ -175,6 +175,10 @@ Section Primitives. (withNextPc (word.add initialL.(getNextPc) (word.of_Z 4)) initialL)) -> mcomp_sat endCycleNormal initialL post; + + spec_leakEvent: forall (initialL: RiscvMachine) (post: unit -> RiscvMachine -> Prop) (e : LeakageEvent), + post tt (withLeakageEvent e initialL) -> + mcomp_sat (leakEvent e) initialL post; }. End Primitives. From f96904a318d4c4507981b42e2b44e6cf666c8c2e Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Fri, 2 Aug 2024 14:39:18 -0400 Subject: [PATCH 24/42] move leakage stuff out of Decode.v --- src/riscv/Spec/Decode.v | 418 +------------------------------- src/riscv/Spec/LeakageOfInstr.v | 416 +++++++++++++++++++++++++++++++ 2 files changed, 419 insertions(+), 415 deletions(-) create mode 100644 src/riscv/Spec/LeakageOfInstr.v diff --git a/src/riscv/Spec/Decode.v b/src/riscv/Spec/Decode.v index 2a28dcc..4bbf55e 100644 --- a/src/riscv/Spec/Decode.v +++ b/src/riscv/Spec/Decode.v @@ -25,12 +25,7 @@ Notation Opcode := BinInt.Z (only parsing). Require Coq.Init.Datatypes. Require Coq.Lists.List. Require Import Coq.ZArith.BinInt. -Require Import Utility.Utility. - -Require Import riscv.Utility.Monads. -Require Import riscv.Utility.MonadNotations. -Section WithMonad. - Context {M : Type -> Type} {MM : Monad M}. +Require Utility.Utility. (* Converted type declarations: *) @@ -60,27 +55,6 @@ Inductive InstructionM64 : Type := | Remuw (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM64 | InvalidM64 : InstructionM64. -Inductive LeakageM64 - {width} {BW : Bitwidth width} {word: word.word width} : Type := -| Mulw_leakage -| Divw_leakage (num : word) (den : word) (*but that this is not.*) -| Divuw_leakage (num : word) (den : word) -| Remw_leakage (num : word) (den : word) (*i'm not sure about this one. probably same as div, so not constant-time.*) -| Remuw_leakage (num : word) (den : word) -| InvalidM64_leakage. - -Definition leakage_of_instr_M64 - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionM64) : M LeakageM64 := - match instr with - | Mulw _ _ _ => Return Mulw_leakage - | Divw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divw_leakage num den) - | Divuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divuw_leakage num den) - | Remw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remw_leakage num den) - | Remuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remuw_leakage num den) - | InvalidM64 => Return InvalidM64_leakage - end. - Inductive InstructionM : Type := | Mul (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM | Mulh (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM @@ -92,33 +66,6 @@ Inductive InstructionM : Type := | Remu (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionM | InvalidM : InstructionM. -Inductive LeakageM - {width} {BW : Bitwidth width} {word: word.word width} : Type := -| Mul_leakage -| Mulh_leakage -| Mulhsu_leakage -| Mulhu_leakage -| Div_leakage (num : word) (den : word) -| Divu_leakage (num : word) (den : word) -| Rem_leakage (num : word) (den : word) -| Remu_leakage (num : word) (den : word) -| InvalidM_leakage. - -Definition leakage_of_instr_M - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionM) : M LeakageM := - match instr with - | Mul _ _ _ => Return Mul_leakage - | Mulh _ _ _ => Return Mulh_leakage - | Mulhsu _ _ _ => Return Mulhsu_leakage - | Mulhu _ _ _ => Return Mulhu_leakage - | Div _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Div_leakage num den) - | Divu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divu_leakage num den) - | Rem _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Rem_leakage num den) - | Remu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remu_leakage num den) - | InvalidM => Return InvalidM_leakage - end. - Inductive InstructionI64 : Type := | Ld (rd : Register) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) : InstructionI64 @@ -138,86 +85,6 @@ Inductive InstructionI64 : Type := | Sraw (rd : Register) (rs1 : Register) (rs2 : Register) : InstructionI64 | InvalidI64 : InstructionI64. -Inductive LeakageI64 -{width} {BW : Bitwidth width} {word: word.word width} - : Type := -| Ld_leakage (addr: word) -| Lwu_leakage (addr: word) -| Addiw_leakage -| Slliw_leakage (shamt : Z) -| Srliw_leakage (shamt : Z) -| Sraiw_leakage (shamt : Z) -| Sd_leakage (addr: word) -| Addw_leakage -| Subw_leakage -| Sllw_leakage (shamt : word) -| Srlw_leakage (shamt : word) -| Sraw_leakage (shamt : word) -| InvalidI64_leakage. - -Definition leakage_of_instr_I64 - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionI64) : M LeakageI64 := - match instr with - | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (word.add rs1_val (word.of_Z oimm12)))) - | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (word.add rs1_val (word.of_Z oimm12)))) - | Addiw _ _ _ => Return Addiw_leakage - | Slliw _ _ shamt => Return (Slliw_leakage shamt) - | Srliw _ _ shamt => Return (Srliw_leakage shamt) - | Sraiw _ _ shamt => Return (Sraiw_leakage shamt) - | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (word.add rs1_val (word.of_Z simm12)))) - | Addw _ _ _ => Return Addw_leakage - | Subw _ _ _ => Return Subw_leakage - | Sllw _ _ rs2 => shamt <- getRegister rs2; Return (Sllw_leakage shamt) - | Srlw _ _ rs2 => shamt <- getRegister rs2; Return (Srlw_leakage shamt) - | Sraw _ _ rs2 => shamt <- getRegister rs2; Return (Sraw_leakage shamt) - | InvalidI64 => Return InvalidI64_leakage - end. - -Inductive LeakageI - {width} {BW : Bitwidth width} {word: word.word width} - : Type := -| Lb_leakage (addr: word) -| Lh_leakage (addr: word) -| Lw_leakage (addr: word) -| Lbu_leakage (addr: word) -| Lhu_leakage (addr: word) -| Fence_leakage (* unsure about this one. *) -| Fence_i_leakage -| Addi_leakage -| Slli_leakage (shamt : Z) -| Slti_leakage -| Sltiu_leakage -| Xori_leakage -| Ori_leakage -| Andi_leakage -| Srli_leakage (shamt : Z) -| Srai_leakage (shamt : Z) -| Auipc_leakage -| Sb_leakage (addr: word) -| Sh_leakage (addr: word) -| Sw_leakage (addr: word) -| Add_leakage -| Sub_leakage -| Sll_leakage (shamt : word) -| Slt_leakage -| Sltu_leakage -| Xor_leakage -| Srl_leakage (shamt : word) -| Sra_leakage (shamt : word) -| Or_leakage -| And_leakage -| Lui_leakage -| Beq_leakage (branch: bool) -| Bne_leakage (branch: bool) -| Blt_leakage (branch: bool) -| Bge_leakage (branch: bool) -| Bltu_leakage (branch: bool) -| Bgeu_leakage (branch: bool) -| Jalr_leakage (* unsure whether i should add the location we're jumping to here - or on the branches. i think not? *) -| Jal_leakage -| InvalidI_leakage. - Inductive InstructionI : Type := | Lb (rd : Register) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) : InstructionI @@ -282,53 +149,6 @@ Inductive InstructionI : Type := | Jal (rd : Register) (jimm20 : Utility.Utility.MachineInt) : InstructionI | InvalidI : InstructionI. -(* are the immediates already signed, or do I need to do some sort of sign extension thing? surely not, since they're already Z, right? *) -Definition leakage_of_instr_I - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionI) : M LeakageI := - match instr with - | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (word.add rs1_val (word.of_Z oimm12))) - | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (word.add rs1_val (word.of_Z oimm12))) - | Lw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lw_leakage (word.add rs1_val (word.of_Z oimm12))) - | Lbu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lbu_leakage (word.add rs1_val (word.of_Z oimm12))) - | Lhu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lhu_leakage (word.add rs1_val (word.of_Z oimm12))) - | Fence _ _ => Return Fence_leakage - | Fence_i => Return Fence_i_leakage - | Addi _ _ _ => Return Addi_leakage - | Slli _ _ shamt => Return (Slli_leakage shamt) - | Slti _ _ _ => Return Slti_leakage - | Sltiu _ _ _ => Return Sltiu_leakage - | Xori _ _ _ => Return Xori_leakage - | Ori _ _ _ => Return Ori_leakage - | Andi _ _ _ => Return Andi_leakage - | Srli _ _ shamt => Return (Srli_leakage shamt) - | Srai _ _ shamt => Return (Srai_leakage shamt) - | Auipc _ _ => Return Auipc_leakage - | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (word.add rs1_val (word.of_Z simm12))) - | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (word.add rs1_val (word.of_Z simm12))) - | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (word.add rs1_val (word.of_Z simm12))) - | Add _ _ _ => Return Add_leakage - | Sub _ _ _ => Return Sub_leakage - | Sll _ _ rs2 => shamt <- getRegister rs2; Return (Sll_leakage shamt) - | Slt _ _ _ => Return Slt_leakage - | Sltu _ _ _ => Return Sltu_leakage - | Xor _ _ _ => Return Xor_leakage - | Srl _ _ rs2 => shamt <- getRegister rs2; Return (Srl_leakage shamt) - | Sra _ _ rs2 => shamt <- getRegister rs2; Return (Sra_leakage shamt) - | Or _ _ _ => Return Or_leakage - | And _ _ _ => Return And_leakage - | Lui _ _ => Return Lui_leakage - | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Beq_leakage (word.eqb rs1_val rs2_val)) - | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bne_leakage (negb (word.eqb rs1_val rs2_val))) - | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Blt_leakage (word.lts rs1_val rs2_val)) - | Bge rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bge_leakage (negb (word.lts rs1_val rs2_val))) - | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bltu_leakage (word.ltu rs1_val rs2_val)) - | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bgeu_leakage (negb (word.ltu rs1_val rs2_val))) - | Jalr _ _ _ => Return Jalr_leakage - | Jal _ _ => Return Jal_leakage - | InvalidI => Return InvalidI_leakage - end. - Inductive InstructionF64 : Type := | Fcvt_l_s (rd : Register) (rs1 : FPRegister) (rm : RoundMode) : InstructionF64 | Fcvt_lu_s (rd : Register) (rs1 : FPRegister) (rm : RoundMode) : InstructionF64 @@ -336,53 +156,6 @@ Inductive InstructionF64 : Type := | Fcvt_s_lu (rd : FPRegister) (rs1 : Register) (rm : RoundMode) : InstructionF64 | InvalidF64 : InstructionF64. -Inductive LeakageF64 : Type := -| Fcvt_l_s_leakage -| Fcvt_lu_s_leakage -| Fcvt_s_l_leakage -| Fcvt_s_lu_leakage -| InvalidF64_leakage. - -Definition leakage_of_instr_F64 (instr : InstructionF64) : M LeakageF64 := - match instr with - | Fcvt_l_s _ _ _ => Return Fcvt_l_s_leakage - | Fcvt_lu_s _ _ _ => Return Fcvt_lu_s_leakage - | Fcvt_s_l _ _ _ => Return Fcvt_s_l_leakage - | Fcvt_s_lu _ _ _ => Return Fcvt_s_lu_leakage - | InvalidF64 => Return InvalidF64_leakage - end. - -Inductive LeakageF - {width} {BW : Bitwidth width} {word: word.word width} - : Type := -| Flw_leakage (addr: word) -| Fsw_leakage (addr: word) -| Fmadd_s_leakage -| Fmsub_s_leakage -| Fnmsub_s_leakage -| Fnmadd_s_leakage -| Fadd_s_leakage -| Fsub_s_leakage -| Fmul_s_leakage -| Fdiv_s_leakage -| Fsqrt_s_leakage -| Fsgnj_s_leakage -| Fsgnjn_s_leakage -| Fsgnjx_s_leakage -| Fmin_s_leakage -| Fmax_s_leakage -| Fcvt_w_s_leakage -| Fcvt_wu_s_leakage -| Fmv_x_w_leakage -| Feq_s_leakage -| Flt_s_leakage -| Fle_s_leakage -| Fclass_s_leakage -| Fcvt_s_w_leakage -| Fcvt_s_wu_leakage -| Fmv_w_x_leakage -| InvalidF_leakage. - Inductive InstructionF : Type := | Flw (rd : FPRegister) (rs1 : Register) (oimm12 : Utility.Utility.MachineInt) : InstructionF @@ -430,59 +203,8 @@ Inductive InstructionF : Type := | Fcvt_s_w (rd : FPRegister) (rs1 : Register) (rm : RoundMode) : InstructionF | Fcvt_s_wu (rd : FPRegister) (rs1 : Register) (rm : RoundMode) : InstructionF | Fmv_w_x (rd : FPRegister) (rs1 : Register) : InstructionF - | InvalidF : InstructionF. Print word.word. - -Definition leakage_of_instr_F - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionF) : M LeakageF := - match instr with - | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (word.add rs1_val (word.of_Z oimm12))) - | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (word.add rs1_val (word.of_Z simm12))) - | Fmadd_s _ _ _ _ _ => Return Fmadd_s_leakage - | Fmsub_s _ _ _ _ _ => Return Fmsub_s_leakage - | Fnmsub_s _ _ _ _ _ => Return Fnmsub_s_leakage - | Fnmadd_s _ _ _ _ _ => Return Fnmadd_s_leakage - | Fadd_s _ _ _ _ => Return Fadd_s_leakage - | Fsub_s _ _ _ _ => Return Fsub_s_leakage - | Fmul_s _ _ _ _ => Return Fmul_s_leakage - | Fdiv_s _ _ _ _ => Return Fdiv_s_leakage - | Fsqrt_s _ _ _ => Return Fsqrt_s_leakage - | Fsgnj_s _ _ _ => Return Fsgnj_s_leakage - | Fsgnjn_s _ _ _ => Return Fsgnjn_s_leakage - | Fsgnjx_s _ _ _ => Return Fsgnjx_s_leakage - | Fmin_s _ _ _ => Return Fmin_s_leakage - | Fmax_s _ _ _ => Return Fmax_s_leakage - | Fcvt_w_s _ _ _ => Return Fcvt_w_s_leakage - | Fcvt_wu_s _ _ _ => Return Fcvt_wu_s_leakage - | Fmv_x_w _ _ => Return Fmv_x_w_leakage - | Feq_s _ _ _ => Return Feq_s_leakage - | Flt_s _ _ _ => Return Flt_s_leakage - | Fle_s _ _ _ => Return Fle_s_leakage - | Fclass_s _ _ => Return Fclass_s_leakage - | Fcvt_s_w _ _ _ => Return Fcvt_s_w_leakage - | Fcvt_s_wu _ _ _ => Return Fcvt_s_wu_leakage - | Fmv_w_x _ _ => Return Fmv_w_x_leakage - | InvalidF => Return InvalidF_leakage - end. - -(* unsure what some of these mean, so unsure whether I have the right definition here. *) -Inductive LeakageCSR : Type := -| Ecall_leakage -| Ebreak_leakage -| Uret_leakage -| Sret_leakage -| Mret_leakage -| Wfi_leakage -| Sfence_vma_leakage -| Csrrw_leakage -| Csrrs_leakage -| Csrrc_leakage -| Csrrwi_leakage -| Csrrsi_leakage -| Csrrci_leakage -| InvalidCSR_leakage. - -(* what are these, and why is it hard to find documentation for them? *) + | InvalidF : InstructionF. + Inductive InstructionCSR : Type := | Ecall : InstructionCSR | Ebreak : InstructionCSR @@ -508,41 +230,6 @@ Inductive InstructionCSR : Type := : InstructionCSR | InvalidCSR : InstructionCSR. -Definition leakage_of_instr_CSR (instr : InstructionCSR) : M LeakageCSR := - match instr with - | Ecall => Return Ecall_leakage - | Ebreak => Return Ebreak_leakage - | Uret => Return Uret_leakage - | Sret => Return Uret_leakage - | Mret => Return Mret_leakage - | Wfi => Return Wfi_leakage - | Sfence_vma _ _ => Return Sfence_vma_leakage - | Csrrw _ _ _ => Return Csrrw_leakage - | Csrrs _ _ _ => Return Csrrs_leakage - | Csrrc _ _ _ => Return Csrrc_leakage - | Csrrwi _ _ _ => Return Csrrwi_leakage - | Csrrsi _ _ _ => Return Csrrsi_leakage - | Csrrci _ _ _ => Return Csrrci_leakage - | InvalidCSR => Return InvalidCSR_leakage - end. - -(* do we care about aqrl here? *) -Inductive LeakageA64 - {width} {BW : Bitwidth width} {word: word.word width} - : Type := -| Lr_d_leakage (addr : word) -| Sc_d_leakage (addr : word) (* behavior of this depends on whether there is a reservation on addr... *) -| Amoswap_d_leakage (addr : word) -| Amoadd_d_leakage (addr : word) -| Amoand_d_leakage (addr : word) -| Amoor_d_leakage (addr : word) -| Amoxor_d_leakage (addr : word) -| Amomax_d_leakage (addr : word) -| Amomaxu_d_leakage (addr : word) -| Amomin_d_leakage (addr : word) -| Amominu_d_leakage (addr : word) -| InvalidA64_leakage. - Inductive InstructionA64 : Type := | Lr_d (rd : Register) (rs1 : Register) (aqrl : Utility.Utility.MachineInt) : InstructionA64 @@ -578,40 +265,6 @@ Inductive InstructionA64 : Type := : InstructionA64 | InvalidA64 : InstructionA64. -Definition leakage_of_instr_A64 - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionA64) : M LeakageA64 := - match instr with - | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage rs1_val) - | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage rs1_val) - | Amoswap_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_d_leakage rs1_val) - | Amoadd_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_d_leakage rs1_val) - | Amoand_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_d_leakage rs1_val) - | Amoor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_d_leakage rs1_val) - | Amoxor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_d_leakage rs1_val) - | Amomax_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_d_leakage rs1_val) - | Amomaxu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_d_leakage rs1_val) - | Amomin_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_d_leakage rs1_val) - | Amominu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_d_leakage rs1_val) - | InvalidA64 => Return InvalidA64_leakage - end. - - Inductive LeakageA - {width} {BW : Bitwidth width} {word: word.word width} - : Type := -| Lr_w_leakage (addr : word) -| Sc_w_leakage (addr : word) -| Amoswap_w_leakage (addr : word) -| Amoadd_w_leakage (addr : word) -| Amoand_w_leakage (addr : word) -| Amoor_w_leakage (addr : word) -| Amoxor_w_leakage (addr : word) -| Amomax_w_leakage (addr : word) -| Amomaxu_w_leakage (addr : word) -| Amomin_w_leakage (addr : word) -| Amominu_w_leakage (addr : word) -| InvalidA_leakage. - Inductive InstructionA : Type := | Lr_w (rd : Register) (rs1 : Register) (aqrl : Utility.Utility.MachineInt) : InstructionA @@ -647,24 +300,6 @@ Inductive InstructionA : Type := : InstructionA | InvalidA : InstructionA. -Definition leakage_of_instr_A - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionA) : M LeakageA := - match instr with - | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage rs1_val) - | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage rs1_val) - | Amoswap_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_w_leakage rs1_val) - | Amoadd_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_w_leakage rs1_val) - | Amoand_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_w_leakage rs1_val) - | Amoor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_w_leakage rs1_val) - | Amoxor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_w_leakage rs1_val) - | Amomax_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_w_leakage rs1_val) - | Amomaxu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_w_leakage rs1_val) - | Amomin_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_w_leakage rs1_val) - | Amominu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_w_leakage rs1_val) - | InvalidA => Return InvalidA_leakage - end. - Inductive Instruction : Type := | IInstruction (iInstruction : InstructionI) : Instruction | MInstruction (mInstruction : InstructionM) : Instruction @@ -677,38 +312,6 @@ Inductive Instruction : Type := | CSRInstruction (csrInstruction : InstructionCSR) : Instruction | InvalidInstruction (inst : Utility.Utility.MachineInt) : Instruction. -Inductive LeakageEvent - {width} {BW : Bitwidth width} {word: word.word width} - : Type := -| ILeakage (iLeakage : LeakageI) -| MLeakage (mLeakage : LeakageM) -| ALeakage (aLeakage : LeakageA) -| FLeakage (fLeakage : LeakageF) -| I64Leakage (i64Leakage : LeakageI64) -| M64Leakage (m64Leakage : LeakageM64) -| A64Leakage (a64Leakage : LeakageA64) -| F64Leakage (f64Leakage : LeakageF64) -| CSRLeakage (csrLeakage : LeakageCSR) -| InvalidLeakage. - -Definition leakage_of_instr - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : Instruction) : M LeakageEvent := - match instr with - | IInstruction instr => l <- leakage_of_instr_I getRegister instr; Return (ILeakage l) - | MInstruction instr => l <- leakage_of_instr_M getRegister instr; Return (MLeakage l) - | AInstruction instr => l <- leakage_of_instr_A getRegister instr; Return (ALeakage l) - | FInstruction instr => l <- leakage_of_instr_F getRegister instr; Return (FLeakage l) - | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; Return (I64Leakage l) - | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; Return (M64Leakage l) - | A64Instruction instr => l <- leakage_of_instr_A64 getRegister instr; Return (A64Leakage l) - | F64Instruction instr => l <- leakage_of_instr_F64 instr; Return (F64Leakage l) - | CSRInstruction instr => l <- leakage_of_instr_CSR instr; Return (CSRLeakage l) - | InvalidInstruction _ => Return InvalidLeakage -end. - -End WithMonad. - (* Converted value declarations: *) (* Skipping instance `Spec.Decode.Eq___InstructionSet' of class @@ -1932,18 +1535,3 @@ Definition decode Utility.Utility.bitSlice Utility.Utility.machineIntToShamt Utility.Utility.signExtend *) - -Definition Mtriv (x : Type) := x. -Definition trivialBind (A B : Type) (x : Mtriv A) (f : A -> B) : Mtriv B := f x. -Definition trivialReturn (A : Type) (a : A) : Mtriv A := a. -Print Monad. -Lemma trivial_left_identity : forall (A B : Type) (a : A) (f : A -> Mtriv B), trivialBind A B (trivialReturn A a) f = f a. -Proof. trivial. Qed. -Lemma trivial_right_identity : forall (A : Type) (m : Mtriv A), trivialBind A A m (trivialReturn A) = m. -Proof. trivial. Qed. -Lemma trivial_associativity : forall (A B C : Type) (m : Mtriv A) (f : A -> Mtriv B) (g : B -> Mtriv C), trivialBind B C (trivialBind A B m f) g = trivialBind A C m (fun x : A => trivialBind B C (f x) g). -Proof. trivial. Qed. -Check @leakage_of_instr. -Definition trivialMonad : Monad Mtriv := - {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. -Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := @leakage_of_instr Mtriv trivialMonad width BW word. diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v new file mode 100644 index 0000000..b95221e --- /dev/null +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -0,0 +1,416 @@ +Require Import Utility.Utility. + +Require Import riscv.Utility.Monads. +Require Import riscv.Utility.MonadNotations. +Require Import riscv.Spec.Decode. + +Notation Register := BinInt.Z (only parsing). + +Section WithMonad. + Context {M : Type -> Type} {MM : Monad M}. + + Inductive LeakageM64 {width} {BW : Bitwidth width} {word: word.word width} : Type := + | Mulw_leakage + | Divw_leakage (num : word) (den : word) + | Divuw_leakage (num : word) (den : word) + | Remw_leakage (num : word) (den : word) + | Remuw_leakage (num : word) (den : word) + | InvalidM64_leakage. + + Definition leakage_of_instr_M64 + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionM64) : M LeakageM64 := + match instr with + | Mulw _ _ _ => Return Mulw_leakage + | Divw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divw_leakage num den) + | Divuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divuw_leakage num den) + | Remw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remw_leakage num den) + | Remuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remuw_leakage num den) + | InvalidM64 => Return InvalidM64_leakage + end. + + Inductive LeakageM + {width} {BW : Bitwidth width} {word: word.word width} : Type := + | Mul_leakage + | Mulh_leakage + | Mulhsu_leakage + | Mulhu_leakage + | Div_leakage (num : word) (den : word) + | Divu_leakage (num : word) (den : word) + | Rem_leakage (num : word) (den : word) + | Remu_leakage (num : word) (den : word) + | InvalidM_leakage. + + Definition leakage_of_instr_M + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionM) : M LeakageM := + match instr with + | Mul _ _ _ => Return Mul_leakage + | Mulh _ _ _ => Return Mulh_leakage + | Mulhsu _ _ _ => Return Mulhsu_leakage + | Mulhu _ _ _ => Return Mulhu_leakage + | Div _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Div_leakage num den) + | Divu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divu_leakage num den) + | Rem _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Rem_leakage num den) + | Remu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remu_leakage num den) + | InvalidM => Return InvalidM_leakage + end. + + Inductive LeakageI64 + {width} {BW : Bitwidth width} {word: word.word width} + : Type := + | Ld_leakage (addr: word) + | Lwu_leakage (addr: word) + | Addiw_leakage + | Slliw_leakage (shamt : Z) + | Srliw_leakage (shamt : Z) + | Sraiw_leakage (shamt : Z) + | Sd_leakage (addr: word) + | Addw_leakage + | Subw_leakage + | Sllw_leakage (shamt : word) + | Srlw_leakage (shamt : word) + | Sraw_leakage (shamt : word) + | InvalidI64_leakage. + + Definition leakage_of_instr_I64 + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionI64) : M LeakageI64 := + match instr with + | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (word.add rs1_val (word.of_Z oimm12)))) + | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (word.add rs1_val (word.of_Z oimm12)))) + | Addiw _ _ _ => Return Addiw_leakage + | Slliw _ _ shamt => Return (Slliw_leakage shamt) + | Srliw _ _ shamt => Return (Srliw_leakage shamt) + | Sraiw _ _ shamt => Return (Sraiw_leakage shamt) + | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (word.add rs1_val (word.of_Z simm12)))) + | Addw _ _ _ => Return Addw_leakage + | Subw _ _ _ => Return Subw_leakage + | Sllw _ _ rs2 => shamt <- getRegister rs2; Return (Sllw_leakage shamt) + | Srlw _ _ rs2 => shamt <- getRegister rs2; Return (Srlw_leakage shamt) + | Sraw _ _ rs2 => shamt <- getRegister rs2; Return (Sraw_leakage shamt) + | InvalidI64 => Return InvalidI64_leakage + end. + + Inductive LeakageI + {width} {BW : Bitwidth width} {word: word.word width} + : Type := + | Lb_leakage (addr: word) + | Lh_leakage (addr: word) + | Lw_leakage (addr: word) + | Lbu_leakage (addr: word) + | Lhu_leakage (addr: word) + | Fence_leakage (* unsure about this one. *) + | Fence_i_leakage + | Addi_leakage + | Slli_leakage (shamt : Z) + | Slti_leakage + | Sltiu_leakage + | Xori_leakage + | Ori_leakage + | Andi_leakage + | Srli_leakage (shamt : Z) + | Srai_leakage (shamt : Z) + | Auipc_leakage + | Sb_leakage (addr: word) + | Sh_leakage (addr: word) + | Sw_leakage (addr: word) + | Add_leakage + | Sub_leakage + | Sll_leakage (shamt : word) + | Slt_leakage + | Sltu_leakage + | Xor_leakage + | Srl_leakage (shamt : word) + | Sra_leakage (shamt : word) + | Or_leakage + | And_leakage + | Lui_leakage + | Beq_leakage (branch: bool) + | Bne_leakage (branch: bool) + | Blt_leakage (branch: bool) + | Bge_leakage (branch: bool) + | Bltu_leakage (branch: bool) + | Bgeu_leakage (branch: bool) + | Jalr_leakage + | Jal_leakage + | InvalidI_leakage. + + Definition leakage_of_instr_I + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionI) : M LeakageI := + match instr with + | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (word.add rs1_val (word.of_Z oimm12))) + | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (word.add rs1_val (word.of_Z oimm12))) + | Lw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lw_leakage (word.add rs1_val (word.of_Z oimm12))) + | Lbu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lbu_leakage (word.add rs1_val (word.of_Z oimm12))) + | Lhu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lhu_leakage (word.add rs1_val (word.of_Z oimm12))) + | Fence _ _ => Return Fence_leakage + | Fence_i => Return Fence_i_leakage + | Addi _ _ _ => Return Addi_leakage + | Slli _ _ shamt => Return (Slli_leakage shamt) + | Slti _ _ _ => Return Slti_leakage + | Sltiu _ _ _ => Return Sltiu_leakage + | Xori _ _ _ => Return Xori_leakage + | Ori _ _ _ => Return Ori_leakage + | Andi _ _ _ => Return Andi_leakage + | Srli _ _ shamt => Return (Srli_leakage shamt) + | Srai _ _ shamt => Return (Srai_leakage shamt) + | Auipc _ _ => Return Auipc_leakage + | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (word.add rs1_val (word.of_Z simm12))) + | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (word.add rs1_val (word.of_Z simm12))) + | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (word.add rs1_val (word.of_Z simm12))) + | Add _ _ _ => Return Add_leakage + | Sub _ _ _ => Return Sub_leakage + | Sll _ _ rs2 => shamt <- getRegister rs2; Return (Sll_leakage shamt) + | Slt _ _ _ => Return Slt_leakage + | Sltu _ _ _ => Return Sltu_leakage + | Xor _ _ _ => Return Xor_leakage + | Srl _ _ rs2 => shamt <- getRegister rs2; Return (Srl_leakage shamt) + | Sra _ _ rs2 => shamt <- getRegister rs2; Return (Sra_leakage shamt) + | Or _ _ _ => Return Or_leakage + | And _ _ _ => Return And_leakage + | Lui _ _ => Return Lui_leakage + | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Beq_leakage (word.eqb rs1_val rs2_val)) + | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bne_leakage (negb (word.eqb rs1_val rs2_val))) + | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Blt_leakage (word.lts rs1_val rs2_val)) + | Bge rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bge_leakage (negb (word.lts rs1_val rs2_val))) + | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bltu_leakage (word.ltu rs1_val rs2_val)) + | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bgeu_leakage (negb (word.ltu rs1_val rs2_val))) + | Jalr _ _ _ => Return Jalr_leakage + | Jal _ _ => Return Jal_leakage + | InvalidI => Return InvalidI_leakage + end. + + Inductive LeakageF64 : Type := + | Fcvt_l_s_leakage + | Fcvt_lu_s_leakage + | Fcvt_s_l_leakage + | Fcvt_s_lu_leakage + | InvalidF64_leakage. + + Definition leakage_of_instr_F64 (instr : InstructionF64) : M LeakageF64 := + match instr with + | Fcvt_l_s _ _ _ => Return Fcvt_l_s_leakage + | Fcvt_lu_s _ _ _ => Return Fcvt_lu_s_leakage + | Fcvt_s_l _ _ _ => Return Fcvt_s_l_leakage + | Fcvt_s_lu _ _ _ => Return Fcvt_s_lu_leakage + | InvalidF64 => Return InvalidF64_leakage + end. + + Inductive LeakageF + {width} {BW : Bitwidth width} {word: word.word width} + : Type := + | Flw_leakage (addr: word) + | Fsw_leakage (addr: word) + | Fmadd_s_leakage + | Fmsub_s_leakage + | Fnmsub_s_leakage + | Fnmadd_s_leakage + | Fadd_s_leakage + | Fsub_s_leakage + | Fmul_s_leakage + | Fdiv_s_leakage + | Fsqrt_s_leakage + | Fsgnj_s_leakage + | Fsgnjn_s_leakage + | Fsgnjx_s_leakage + | Fmin_s_leakage + | Fmax_s_leakage + | Fcvt_w_s_leakage + | Fcvt_wu_s_leakage + | Fmv_x_w_leakage + | Feq_s_leakage + | Flt_s_leakage + | Fle_s_leakage + | Fclass_s_leakage + | Fcvt_s_w_leakage + | Fcvt_s_wu_leakage + | Fmv_w_x_leakage + | InvalidF_leakage. + + Definition leakage_of_instr_F + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionF) : M LeakageF := + match instr with + | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (word.add rs1_val (word.of_Z oimm12))) + | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (word.add rs1_val (word.of_Z simm12))) + | Fmadd_s _ _ _ _ _ => Return Fmadd_s_leakage + | Fmsub_s _ _ _ _ _ => Return Fmsub_s_leakage + | Fnmsub_s _ _ _ _ _ => Return Fnmsub_s_leakage + | Fnmadd_s _ _ _ _ _ => Return Fnmadd_s_leakage + | Fadd_s _ _ _ _ => Return Fadd_s_leakage + | Fsub_s _ _ _ _ => Return Fsub_s_leakage + | Fmul_s _ _ _ _ => Return Fmul_s_leakage + | Fdiv_s _ _ _ _ => Return Fdiv_s_leakage + | Fsqrt_s _ _ _ => Return Fsqrt_s_leakage + | Fsgnj_s _ _ _ => Return Fsgnj_s_leakage + | Fsgnjn_s _ _ _ => Return Fsgnjn_s_leakage + | Fsgnjx_s _ _ _ => Return Fsgnjx_s_leakage + | Fmin_s _ _ _ => Return Fmin_s_leakage + | Fmax_s _ _ _ => Return Fmax_s_leakage + | Fcvt_w_s _ _ _ => Return Fcvt_w_s_leakage + | Fcvt_wu_s _ _ _ => Return Fcvt_wu_s_leakage + | Fmv_x_w _ _ => Return Fmv_x_w_leakage + | Feq_s _ _ _ => Return Feq_s_leakage + | Flt_s _ _ _ => Return Flt_s_leakage + | Fle_s _ _ _ => Return Fle_s_leakage + | Fclass_s _ _ => Return Fclass_s_leakage + | Fcvt_s_w _ _ _ => Return Fcvt_s_w_leakage + | Fcvt_s_wu _ _ _ => Return Fcvt_s_wu_leakage + | Fmv_w_x _ _ => Return Fmv_w_x_leakage + | InvalidF => Return InvalidF_leakage + end. + + Inductive LeakageCSR : Type := + | Ecall_leakage + | Ebreak_leakage + | Uret_leakage + | Sret_leakage + | Mret_leakage + | Wfi_leakage + | Sfence_vma_leakage + | Csrrw_leakage + | Csrrs_leakage + | Csrrc_leakage + | Csrrwi_leakage + | Csrrsi_leakage + | Csrrci_leakage + | InvalidCSR_leakage. + + Definition leakage_of_instr_CSR (instr : InstructionCSR) : M LeakageCSR := + match instr with + | Ecall => Return Ecall_leakage + | Ebreak => Return Ebreak_leakage + | Uret => Return Uret_leakage + | Sret => Return Uret_leakage + | Mret => Return Mret_leakage + | Wfi => Return Wfi_leakage + | Sfence_vma _ _ => Return Sfence_vma_leakage + | Csrrw _ _ _ => Return Csrrw_leakage + | Csrrs _ _ _ => Return Csrrs_leakage + | Csrrc _ _ _ => Return Csrrc_leakage + | Csrrwi _ _ _ => Return Csrrwi_leakage + | Csrrsi _ _ _ => Return Csrrsi_leakage + | Csrrci _ _ _ => Return Csrrci_leakage + | InvalidCSR => Return InvalidCSR_leakage + end. + + Inductive LeakageA64 + {width} {BW : Bitwidth width} {word: word.word width} + : Type := + | Lr_d_leakage (addr : word) + | Sc_d_leakage (addr : word) + | Amoswap_d_leakage (addr : word) + | Amoadd_d_leakage (addr : word) + | Amoand_d_leakage (addr : word) + | Amoor_d_leakage (addr : word) + | Amoxor_d_leakage (addr : word) + | Amomax_d_leakage (addr : word) + | Amomaxu_d_leakage (addr : word) + | Amomin_d_leakage (addr : word) + | Amominu_d_leakage (addr : word) + | InvalidA64_leakage. + + + Definition leakage_of_instr_A64 + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionA64) : M LeakageA64 := + match instr with + | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage rs1_val) + | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage rs1_val) + | Amoswap_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_d_leakage rs1_val) + | Amoadd_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_d_leakage rs1_val) + | Amoand_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_d_leakage rs1_val) + | Amoor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_d_leakage rs1_val) + | Amoxor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_d_leakage rs1_val) + | Amomax_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_d_leakage rs1_val) + | Amomaxu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_d_leakage rs1_val) + | Amomin_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_d_leakage rs1_val) + | Amominu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_d_leakage rs1_val) + | InvalidA64 => Return InvalidA64_leakage + end. + + Inductive LeakageA + {width} {BW : Bitwidth width} {word: word.word width} + : Type := + | Lr_w_leakage (addr : word) + | Sc_w_leakage (addr : word) + | Amoswap_w_leakage (addr : word) + | Amoadd_w_leakage (addr : word) + | Amoand_w_leakage (addr : word) + | Amoor_w_leakage (addr : word) + | Amoxor_w_leakage (addr : word) + | Amomax_w_leakage (addr : word) + | Amomaxu_w_leakage (addr : word) + | Amomin_w_leakage (addr : word) + | Amominu_w_leakage (addr : word) + | InvalidA_leakage. + + Definition leakage_of_instr_A + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : InstructionA) : M LeakageA := + match instr with + | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage rs1_val) + | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage rs1_val) + | Amoswap_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_w_leakage rs1_val) + | Amoadd_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_w_leakage rs1_val) + | Amoand_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_w_leakage rs1_val) + | Amoor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_w_leakage rs1_val) + | Amoxor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_w_leakage rs1_val) + | Amomax_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_w_leakage rs1_val) + | Amomaxu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_w_leakage rs1_val) + | Amomin_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_w_leakage rs1_val) + | Amominu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_w_leakage rs1_val) + | InvalidA => Return InvalidA_leakage + end. + + Inductive LeakageEvent + {width} {BW : Bitwidth width} {word: word.word width} + : Type := + | ILeakage (iLeakage : LeakageI) + | MLeakage (mLeakage : LeakageM) + | ALeakage (aLeakage : LeakageA) + | FLeakage (fLeakage : LeakageF) + | I64Leakage (i64Leakage : LeakageI64) + | M64Leakage (m64Leakage : LeakageM64) + | A64Leakage (a64Leakage : LeakageA64) + | F64Leakage (f64Leakage : LeakageF64) + | CSRLeakage (csrLeakage : LeakageCSR) + | InvalidLeakage. + + Definition leakage_of_instr + {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) + (instr : Instruction) : M LeakageEvent := + match instr with + | IInstruction instr => l <- leakage_of_instr_I getRegister instr; Return (ILeakage l) + | MInstruction instr => l <- leakage_of_instr_M getRegister instr; Return (MLeakage l) + | AInstruction instr => l <- leakage_of_instr_A getRegister instr; Return (ALeakage l) + | FInstruction instr => l <- leakage_of_instr_F getRegister instr; Return (FLeakage l) + | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; Return (I64Leakage l) + | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; Return (M64Leakage l) + | A64Instruction instr => l <- leakage_of_instr_A64 getRegister instr; Return (A64Leakage l) + | F64Instruction instr => l <- leakage_of_instr_F64 instr; Return (F64Leakage l) + | CSRInstruction instr => l <- leakage_of_instr_CSR instr; Return (CSRLeakage l) + | InvalidInstruction _ => Return InvalidLeakage + end. + +End WithMonad. + +Definition Mtriv (x : Type) := x. +Definition trivialBind (A B : Type) (x : Mtriv A) (f : A -> B) : Mtriv B := f x. +Definition trivialReturn (A : Type) (a : A) : Mtriv A := a. + +Lemma trivial_left_identity : forall (A B : Type) (a : A) (f : A -> Mtriv B), trivialBind A B (trivialReturn A a) f = f a. +Proof. trivial. Qed. + +Lemma trivial_right_identity : forall (A : Type) (m : Mtriv A), trivialBind A A m (trivialReturn A) = m. +Proof. trivial. Qed. + +Lemma trivial_associativity : forall (A B C : Type) (m : Mtriv A) (f : A -> Mtriv B) (g : B -> Mtriv C), trivialBind B C (trivialBind A B m f) g = trivialBind A C m (fun x : A => trivialBind B C (f x) g). +Proof. trivial. Qed. + +Definition trivialMonad : Monad Mtriv := + {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. + +Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := @leakage_of_instr Mtriv trivialMonad width BW word. From fa470cab5d32a2b165d148a97b9f168ed73d5bfe Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 11:18:03 -0400 Subject: [PATCH 25/42] move identity_Monad where it belongs --- src/riscv/Spec/LeakageOfInstr.v | 18 +----------------- src/riscv/Utility/Monads.v | 7 +++++++ 2 files changed, 8 insertions(+), 17 deletions(-) diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index b95221e..d8c1605 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -397,20 +397,4 @@ Section WithMonad. End WithMonad. -Definition Mtriv (x : Type) := x. -Definition trivialBind (A B : Type) (x : Mtriv A) (f : A -> B) : Mtriv B := f x. -Definition trivialReturn (A : Type) (a : A) : Mtriv A := a. - -Lemma trivial_left_identity : forall (A B : Type) (a : A) (f : A -> Mtriv B), trivialBind A B (trivialReturn A a) f = f a. -Proof. trivial. Qed. - -Lemma trivial_right_identity : forall (A : Type) (m : Mtriv A), trivialBind A A m (trivialReturn A) = m. -Proof. trivial. Qed. - -Lemma trivial_associativity : forall (A B C : Type) (m : Mtriv A) (f : A -> Mtriv B) (g : B -> Mtriv C), trivialBind B C (trivialBind A B m f) g = trivialBind A C m (fun x : A => trivialBind B C (f x) g). -Proof. trivial. Qed. - -Definition trivialMonad : Monad Mtriv := - {| Bind := trivialBind; Return := trivialReturn; left_identity := trivial_left_identity; right_identity := trivial_right_identity; associativity := trivial_associativity |}. - -Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := @leakage_of_instr Mtriv trivialMonad width BW word. +Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := @leakage_of_instr (fun T => T) _ width BW word. diff --git a/src/riscv/Utility/Monads.v b/src/riscv/Utility/Monads.v index a4ebc17..8bf739c 100644 --- a/src/riscv/Utility/Monads.v +++ b/src/riscv/Utility/Monads.v @@ -37,6 +37,13 @@ Ltac prove_monad_law := | o: option _ |- _ => destruct o end. +#[global] Instance identity_Monad: Monad (fun T => T). refine ({| + Bind A B (x: A) (f: A -> B) := f x; + Return A (a: A) := a +|}). +all: prove_monad_law. +Defined. + #[global] Instance option_Monad: Monad option. refine ({| Bind A B (o: option A) (f: A -> option B) := match o with From ff149f8cfe80eaf2380e06e66307835fa9cf4e4a Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 11:32:14 -0400 Subject: [PATCH 26/42] Decided how to write leakage spec when idk what instruction should leak. Just use option type. When the thing returns None, semantics require proving postcondition about appending 'arbitrary garbage' to leakage trace. Arguably it would make more sense to say that the machine just gets stuck when the thing returns None, but that could be inconvenient if somebody wants to prove things about some instructions without having to define what their leakage is... --- src/riscv/Spec/LeakageOfInstr.v | 14 ++++++++++---- 1 file changed, 10 insertions(+), 4 deletions(-) diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index d8c1605..e51900d 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -262,7 +262,7 @@ Section WithMonad. | InvalidF => Return InvalidF_leakage end. - Inductive LeakageCSR : Type := + (*Inductive LeakageCSR : Type := | Ecall_leakage | Ebreak_leakage | Uret_leakage @@ -276,9 +276,11 @@ Section WithMonad. | Csrrwi_leakage | Csrrsi_leakage | Csrrci_leakage - | InvalidCSR_leakage. + | InvalidCSR_leakage.*) - Definition leakage_of_instr_CSR (instr : InstructionCSR) : M LeakageCSR := + Inductive LeakageCSR : Type :=. + + (*Definition leakage_of_instr_CSR (instr : InstructionCSR) : M LeakageCSR := match instr with | Ecall => Return Ecall_leakage | Ebreak => Return Ebreak_leakage @@ -294,7 +296,10 @@ Section WithMonad. | Csrrsi _ _ _ => Return Csrrsi_leakage | Csrrci _ _ _ => Return Csrrci_leakage | InvalidCSR => Return InvalidCSR_leakage - end. + end.*) + + Definition leakage_of_instr_CSR (instr : InstructionCSR) : M (option LeakageCSR) := + Return None. Inductive LeakageA64 {width} {BW : Bitwidth width} {word: word.word width} @@ -368,6 +373,7 @@ Section WithMonad. Inductive LeakageEvent {width} {BW : Bitwidth width} {word: word.word width} : Type := + | arbitrary_garbage (n : nat) | ILeakage (iLeakage : LeakageI) | MLeakage (mLeakage : LeakageM) | ALeakage (aLeakage : LeakageA) From ab888f9ef193bbce38982998e31204500496e8e7 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 20:11:47 -0400 Subject: [PATCH 27/42] fixup LeakageOfInstr to not make claims about instructions i don't understand --- src/riscv/Spec/LeakageOfInstr.v | 202 ++++++-------------------------- 1 file changed, 36 insertions(+), 166 deletions(-) diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index e51900d..f75a7bb 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -17,9 +17,8 @@ Section WithMonad. | Remuw_leakage (num : word) (den : word) | InvalidM64_leakage. - Definition leakage_of_instr_M64 - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionM64) : M LeakageM64 := + Definition leakage_of_instr_M64 {width} {BW : Bitwidth width} {word: word.word width} + (getRegister : Register -> M word) (instr : InstructionM64) : M LeakageM64 := match instr with | Mulw _ _ _ => Return Mulw_leakage | Divw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divw_leakage num den) @@ -29,8 +28,7 @@ Section WithMonad. | InvalidM64 => Return InvalidM64_leakage end. - Inductive LeakageM - {width} {BW : Bitwidth width} {word: word.word width} : Type := + Inductive LeakageM {width} {BW : Bitwidth width} {word: word.word width} : Type := | Mul_leakage | Mulh_leakage | Mulhsu_leakage @@ -41,9 +39,8 @@ Section WithMonad. | Remu_leakage (num : word) (den : word) | InvalidM_leakage. - Definition leakage_of_instr_M - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionM) : M LeakageM := + Definition leakage_of_instr_M {width} {BW : Bitwidth width} {word: word.word width} + (getRegister : Register -> M word) (instr : InstructionM) : M LeakageM := match instr with | Mul _ _ _ => Return Mul_leakage | Mulh _ _ _ => Return Mulh_leakage @@ -56,9 +53,7 @@ Section WithMonad. | InvalidM => Return InvalidM_leakage end. - Inductive LeakageI64 - {width} {BW : Bitwidth width} {word: word.word width} - : Type := + Inductive LeakageI64 {width} {BW : Bitwidth width} {word: word.word width} : Type := | Ld_leakage (addr: word) | Lwu_leakage (addr: word) | Addiw_leakage @@ -73,17 +68,16 @@ Section WithMonad. | Sraw_leakage (shamt : word) | InvalidI64_leakage. - Definition leakage_of_instr_I64 - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionI64) : M LeakageI64 := + Definition leakage_of_instr_I64 {width} {BW : Bitwidth width} {word: word.word width} + (getRegister : Register -> M word) (instr : InstructionI64) : M LeakageI64 := match instr with - | Ld _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Ld_leakage (word.add rs1_val (word.of_Z oimm12)))) - | Lwu _ rs1 oimm12 => Bind (getRegister rs1) (fun rs1_val => Return (Lwu_leakage (word.add rs1_val (word.of_Z oimm12)))) + | Ld _ rs1 oimm12 => addr <- getRegister rs1; Return (Ld_leakage (word.add addr (word.of_Z oimm12))) + | Lwu _ rs1 oimm12 => addr <- getRegister rs1; Return (Lwu_leakage (word.add addr (word.of_Z oimm12))) | Addiw _ _ _ => Return Addiw_leakage | Slliw _ _ shamt => Return (Slliw_leakage shamt) | Srliw _ _ shamt => Return (Srliw_leakage shamt) | Sraiw _ _ shamt => Return (Sraiw_leakage shamt) - | Sd rs1 _ simm12 => Bind (getRegister rs1) (fun rs1_val => Return (Sd_leakage (word.add rs1_val (word.of_Z simm12)))) + | Sd rs1 _ simm12 => addr <- getRegister rs1; Return (Sd_leakage (word.add addr (word.of_Z simm12))) | Addw _ _ _ => Return Addw_leakage | Subw _ _ _ => Return Subw_leakage | Sllw _ _ rs2 => shamt <- getRegister rs2; Return (Sllw_leakage shamt) @@ -92,9 +86,7 @@ Section WithMonad. | InvalidI64 => Return InvalidI64_leakage end. - Inductive LeakageI - {width} {BW : Bitwidth width} {word: word.word width} - : Type := + Inductive LeakageI {width} {BW : Bitwidth width} {word: word.word width} : Type := | Lb_leakage (addr: word) | Lh_leakage (addr: word) | Lw_leakage (addr: word) @@ -136,9 +128,8 @@ Section WithMonad. | Jal_leakage | InvalidI_leakage. - Definition leakage_of_instr_I - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionI) : M LeakageI := + Definition leakage_of_instr_I {width} {BW : Bitwidth width} {word: word.word width} + (getRegister : Register -> M word) (instr : InstructionI) : M LeakageI := match instr with | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (word.add rs1_val (word.of_Z oimm12))) | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (word.add rs1_val (word.of_Z oimm12))) @@ -182,25 +173,9 @@ Section WithMonad. | InvalidI => Return InvalidI_leakage end. - Inductive LeakageF64 : Type := - | Fcvt_l_s_leakage - | Fcvt_lu_s_leakage - | Fcvt_s_l_leakage - | Fcvt_s_lu_leakage - | InvalidF64_leakage. + Inductive LeakageF64 : Type :=. - Definition leakage_of_instr_F64 (instr : InstructionF64) : M LeakageF64 := - match instr with - | Fcvt_l_s _ _ _ => Return Fcvt_l_s_leakage - | Fcvt_lu_s _ _ _ => Return Fcvt_lu_s_leakage - | Fcvt_s_l _ _ _ => Return Fcvt_s_l_leakage - | Fcvt_s_lu _ _ _ => Return Fcvt_s_lu_leakage - | InvalidF64 => Return InvalidF64_leakage - end. - - Inductive LeakageF - {width} {BW : Bitwidth width} {word: word.word width} - : Type := + Inductive LeakageF {width} {BW : Bitwidth width} {word: word.word width} : Type := | Flw_leakage (addr: word) | Fsw_leakage (addr: word) | Fmadd_s_leakage @@ -229,9 +204,8 @@ Section WithMonad. | Fmv_w_x_leakage | InvalidF_leakage. - Definition leakage_of_instr_F - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionF) : M LeakageF := + Definition leakage_of_instr_F {width} {BW : Bitwidth width} {word: word.word width} + (getRegister : Register -> M word) (instr : InstructionF) : M LeakageF := match instr with | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (word.add rs1_val (word.of_Z oimm12))) | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (word.add rs1_val (word.of_Z simm12))) @@ -262,118 +236,14 @@ Section WithMonad. | InvalidF => Return InvalidF_leakage end. - (*Inductive LeakageCSR : Type := - | Ecall_leakage - | Ebreak_leakage - | Uret_leakage - | Sret_leakage - | Mret_leakage - | Wfi_leakage - | Sfence_vma_leakage - | Csrrw_leakage - | Csrrs_leakage - | Csrrc_leakage - | Csrrwi_leakage - | Csrrsi_leakage - | Csrrci_leakage - | InvalidCSR_leakage.*) - Inductive LeakageCSR : Type :=. - - (*Definition leakage_of_instr_CSR (instr : InstructionCSR) : M LeakageCSR := - match instr with - | Ecall => Return Ecall_leakage - | Ebreak => Return Ebreak_leakage - | Uret => Return Uret_leakage - | Sret => Return Uret_leakage - | Mret => Return Mret_leakage - | Wfi => Return Wfi_leakage - | Sfence_vma _ _ => Return Sfence_vma_leakage - | Csrrw _ _ _ => Return Csrrw_leakage - | Csrrs _ _ _ => Return Csrrs_leakage - | Csrrc _ _ _ => Return Csrrc_leakage - | Csrrwi _ _ _ => Return Csrrwi_leakage - | Csrrsi _ _ _ => Return Csrrsi_leakage - | Csrrci _ _ _ => Return Csrrci_leakage - | InvalidCSR => Return InvalidCSR_leakage - end.*) - - Definition leakage_of_instr_CSR (instr : InstructionCSR) : M (option LeakageCSR) := - Return None. - Inductive LeakageA64 - {width} {BW : Bitwidth width} {word: word.word width} - : Type := - | Lr_d_leakage (addr : word) - | Sc_d_leakage (addr : word) - | Amoswap_d_leakage (addr : word) - | Amoadd_d_leakage (addr : word) - | Amoand_d_leakage (addr : word) - | Amoor_d_leakage (addr : word) - | Amoxor_d_leakage (addr : word) - | Amomax_d_leakage (addr : word) - | Amomaxu_d_leakage (addr : word) - | Amomin_d_leakage (addr : word) - | Amominu_d_leakage (addr : word) - | InvalidA64_leakage. + Inductive LeakageA64 : Type :=. - - Definition leakage_of_instr_A64 - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionA64) : M LeakageA64 := - match instr with - | Lr_d _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_d_leakage rs1_val) - | Sc_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_d_leakage rs1_val) - | Amoswap_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_d_leakage rs1_val) - | Amoadd_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_d_leakage rs1_val) - | Amoand_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_d_leakage rs1_val) - | Amoor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_d_leakage rs1_val) - | Amoxor_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_d_leakage rs1_val) - | Amomax_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_d_leakage rs1_val) - | Amomaxu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_d_leakage rs1_val) - | Amomin_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_d_leakage rs1_val) - | Amominu_d _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_d_leakage rs1_val) - | InvalidA64 => Return InvalidA64_leakage - end. - - Inductive LeakageA - {width} {BW : Bitwidth width} {word: word.word width} - : Type := - | Lr_w_leakage (addr : word) - | Sc_w_leakage (addr : word) - | Amoswap_w_leakage (addr : word) - | Amoadd_w_leakage (addr : word) - | Amoand_w_leakage (addr : word) - | Amoor_w_leakage (addr : word) - | Amoxor_w_leakage (addr : word) - | Amomax_w_leakage (addr : word) - | Amomaxu_w_leakage (addr : word) - | Amomin_w_leakage (addr : word) - | Amominu_w_leakage (addr : word) - | InvalidA_leakage. - - Definition leakage_of_instr_A - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : InstructionA) : M LeakageA := - match instr with - | Lr_w _ rs1 _ => rs1_val <- getRegister rs1; Return (Lr_w_leakage rs1_val) - | Sc_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Sc_w_leakage rs1_val) - | Amoswap_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoswap_w_leakage rs1_val) - | Amoadd_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoadd_w_leakage rs1_val) - | Amoand_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoand_w_leakage rs1_val) - | Amoor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoor_w_leakage rs1_val) - | Amoxor_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amoxor_w_leakage rs1_val) - | Amomax_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomax_w_leakage rs1_val) - | Amomaxu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomaxu_w_leakage rs1_val) - | Amomin_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amomin_w_leakage rs1_val) - | Amominu_w _ rs1 _ _ => rs1_val <- getRegister rs1; Return (Amominu_w_leakage rs1_val) - | InvalidA => Return InvalidA_leakage - end. + Inductive LeakageA : Type :=. - Inductive LeakageEvent - {width} {BW : Bitwidth width} {word: word.word width} - : Type := - | arbitrary_garbage (n : nat) + Inductive LeakageEvent {width} {BW : Bitwidth width} {word: word.word width} : Type := + | anything {X : Type} (x : X) | ILeakage (iLeakage : LeakageI) | MLeakage (mLeakage : LeakageM) | ALeakage (aLeakage : LeakageA) @@ -385,22 +255,22 @@ Section WithMonad. | CSRLeakage (csrLeakage : LeakageCSR) | InvalidLeakage. - Definition leakage_of_instr - {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) - (instr : Instruction) : M LeakageEvent := + Definition leakage_of_instr {width} {BW : Bitwidth width} {word: word.word width} + (getRegister : Register -> M word) (instr : Instruction) : M (option LeakageEvent) := match instr with - | IInstruction instr => l <- leakage_of_instr_I getRegister instr; Return (ILeakage l) - | MInstruction instr => l <- leakage_of_instr_M getRegister instr; Return (MLeakage l) - | AInstruction instr => l <- leakage_of_instr_A getRegister instr; Return (ALeakage l) - | FInstruction instr => l <- leakage_of_instr_F getRegister instr; Return (FLeakage l) - | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; Return (I64Leakage l) - | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; Return (M64Leakage l) - | A64Instruction instr => l <- leakage_of_instr_A64 getRegister instr; Return (A64Leakage l) - | F64Instruction instr => l <- leakage_of_instr_F64 instr; Return (F64Leakage l) - | CSRInstruction instr => l <- leakage_of_instr_CSR instr; Return (CSRLeakage l) - | InvalidInstruction _ => Return InvalidLeakage + | IInstruction instr => l <- leakage_of_instr_I getRegister instr; Return (Some (ILeakage l)) + | MInstruction instr => l <- leakage_of_instr_M getRegister instr; Return (Some (MLeakage l)) + | AInstruction instr => Return None + | FInstruction instr => l <- leakage_of_instr_F getRegister instr; Return (Some (FLeakage l)) + | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; Return (Some (I64Leakage l)) + | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; Return (Some (M64Leakage l)) + | A64Instruction instr => Return None + | F64Instruction instr => Return None + | CSRInstruction instr => Return None + | InvalidInstruction _ => Return (Some InvalidLeakage) end. End WithMonad. -Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := @leakage_of_instr (fun T => T) _ width BW word. +Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := + @leakage_of_instr (fun T => T) _ width BW word. From b13861b7b9b67bbadd90827b2aeeba93b2ec25fe Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 20:23:16 -0400 Subject: [PATCH 28/42] adapt other files to leakage_of_instr returning option type --- src/riscv/Platform/MetricRiscvMachine.v | 1 + src/riscv/Platform/RiscvMachine.v | 1 + src/riscv/Platform/Run.v | 1 + src/riscv/Spec/Machine.v | 3 ++- src/riscv/Spec/MetricPrimitives.v | 7 +++++-- src/riscv/Spec/Primitives.v | 7 +++++-- 6 files changed, 15 insertions(+), 5 deletions(-) diff --git a/src/riscv/Platform/MetricRiscvMachine.v b/src/riscv/Platform/MetricRiscvMachine.v index c1b17b7..7a64963 100644 --- a/src/riscv/Platform/MetricRiscvMachine.v +++ b/src/riscv/Platform/MetricRiscvMachine.v @@ -3,6 +3,7 @@ Require Import coqutil.Map.Interface. Require Import coqutil.Word.Interface. Require Import coqutil.Word.LittleEndian. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Platform.Memory. Require Import riscv.Utility.Utility. Require Import riscv.Platform.RiscvMachine. diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index da30e7e..5041e2f 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -4,6 +4,7 @@ Require Import coqutil.Map.Interface. Require Import coqutil.Word.Interface. Require Import coqutil.Word.LittleEndian. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Platform.Memory. Require Import riscv.Utility.Utility. diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index 0623084..94178ee 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -3,6 +3,7 @@ Require Import coqutil.Word.LittleEndian. Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Platform.Memory. (* should go before Program because both define loadByte etc *) Require Import riscv.Spec.Machine. Require Import riscv.Spec.Execute. diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index 96ae1b0..c9fe29a 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -3,6 +3,7 @@ Require Import riscv.Utility.Monads. Require riscv.Utility.MonadNotations. Require Import riscv.Utility.Utility. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Spec.CSRField. Local Open Scope Z_scope. @@ -71,7 +72,7 @@ Class RiscvProgramWithLeakage {M}{t}{MM}{MWt}:= mkRiscvProgramWithLeakage { RVP :> @RiscvProgram M t MM MWt; - leakEvent : @LeakageEvent width BW word -> M unit; + leakEvent : (option LeakageEvent) -> M unit; }. Class RiscvMachine`{MP: RiscvProgram} := mkRiscvMachine { diff --git a/src/riscv/Spec/MetricPrimitives.v b/src/riscv/Spec/MetricPrimitives.v index 8deae7f..927dad6 100644 --- a/src/riscv/Spec/MetricPrimitives.v +++ b/src/riscv/Spec/MetricPrimitives.v @@ -4,6 +4,7 @@ Require Import coqutil.Map.Interface. Require Import riscv.Utility.Monads. Require Import riscv.Utility.Utility. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Platform.Memory. Require Import riscv.Platform.RiscvMachine. Require Import riscv.Spec.Machine. @@ -136,8 +137,10 @@ Section MetricPrimitives. initialL))) -> mcomp_sat endCycleNormal initialL post; - spec_leakEvent: forall (initialL: MetricRiscvMachine) (post: unit -> MetricRiscvMachine -> Prop) (e : LeakageEvent), - post tt (withLeakageEvent e initialL) -> + spec_leakEvent: forall (initialL: MetricRiscvMachine) (post: unit -> MetricRiscvMachine -> Prop) (e : option LeakageEvent), + (match e with + | Some e => post tt (withLeakageEvent e initialL) + | None => forall X (x : X), post tt (withLeakageEvent (anything x) initialL) end) -> mcomp_sat (leakEvent e) initialL post; }. diff --git a/src/riscv/Spec/Primitives.v b/src/riscv/Spec/Primitives.v index 911c7b8..c6ea9e9 100644 --- a/src/riscv/Spec/Primitives.v +++ b/src/riscv/Spec/Primitives.v @@ -4,6 +4,7 @@ Require Import coqutil.Map.Interface. Require Import riscv.Utility.Monads. Require Import riscv.Utility.Utility. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Platform.Memory. Require Import riscv.Spec.Machine. Require Import riscv.Platform.RiscvMachine. @@ -176,8 +177,10 @@ Section Primitives. initialL)) -> mcomp_sat endCycleNormal initialL post; - spec_leakEvent: forall (initialL: RiscvMachine) (post: unit -> RiscvMachine -> Prop) (e : LeakageEvent), - post tt (withLeakageEvent e initialL) -> + spec_leakEvent: forall (initialL: RiscvMachine) (post: unit -> RiscvMachine -> Prop) (e : option LeakageEvent), + (match e with + | Some e => post tt (withLeakageEvent e initialL) + | None => forall X (x : X), post tt (withLeakageEvent (anything x) initialL) end) -> mcomp_sat (leakEvent e) initialL post; }. From 2d146e181159ca2b8a61b1d78369599511cd4fef Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 20:26:26 -0400 Subject: [PATCH 29/42] leak instruction-fetch addressses --- src/riscv/Platform/Run.v | 1 + src/riscv/Spec/LeakageOfInstr.v | 1 + 2 files changed, 2 insertions(+) diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index 94178ee..e777473 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -22,6 +22,7 @@ Section Riscv. Definition run1(iset: InstructionSet): M unit := pc <- getPC; + leakEvent (Some (fetchInstr pc));; inst <- loadWord Fetch pc; let inst' := decode iset (combine 4 inst) in leakage_event <- leakage_of_instr getRegister inst'; diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index f75a7bb..982fc22 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -244,6 +244,7 @@ Section WithMonad. Inductive LeakageEvent {width} {BW : Bitwidth width} {word: word.word width} : Type := | anything {X : Type} (x : X) + | fetchInstr (address : word) | ILeakage (iLeakage : LeakageI) | MLeakage (mLeakage : LeakageM) | ALeakage (aLeakage : LeakageA) From c1d51216e4522e9ca8e4c0891d926bb849a886d5 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 21:57:24 -0400 Subject: [PATCH 30/42] fix remaining errors, so that it compiles --- src/riscv/Platform/MaterializeRiscvProgram.v | 3 ++- .../Platform/MetricMaterializeRiscvProgram.v | 5 +++++ src/riscv/Platform/MetricSane.v | 2 ++ src/riscv/Platform/Minimal.v | 21 +++++++++++++++---- src/riscv/Platform/MinimalMMIO.v | 10 +++++++-- 5 files changed, 34 insertions(+), 7 deletions(-) diff --git a/src/riscv/Platform/MaterializeRiscvProgram.v b/src/riscv/Platform/MaterializeRiscvProgram.v index d1065ae..f71add6 100644 --- a/src/riscv/Platform/MaterializeRiscvProgram.v +++ b/src/riscv/Platform/MaterializeRiscvProgram.v @@ -1,6 +1,7 @@ Require Import riscv.Utility.Monads. Require Import riscv.Utility.FreeMonad. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Spec.Machine. Require Import riscv.Utility.MkMachineWidth. Require Import riscv.Utility.Utility. @@ -27,7 +28,7 @@ Section Riscv. | GetPrivMode | SetPrivMode (_ : PrivMode) | Fence (_ : MachineInt) (_ : MachineInt) - | LeakEvent (_ : LeakageEvent) + | LeakEvent (_ : option LeakageEvent) | GetPC | SetPC (_ : word) | StartCycle diff --git a/src/riscv/Platform/MetricMaterializeRiscvProgram.v b/src/riscv/Platform/MetricMaterializeRiscvProgram.v index 688ed7b..4335be2 100644 --- a/src/riscv/Platform/MetricMaterializeRiscvProgram.v +++ b/src/riscv/Platform/MetricMaterializeRiscvProgram.v @@ -39,4 +39,9 @@ Section Riscv. endCycleEarly A := act (addMetricInstructions 1, EndCycleEarly A) ret; |}. + Global Instance MetricMaterializeWithLeakage : RiscvProgramWithLeakage := {| + RVP := MetricMaterialize; + leakEvent a := act (id, LeakEvent a) ret + |}. + End Riscv. diff --git a/src/riscv/Platform/MetricSane.v b/src/riscv/Platform/MetricSane.v index ccc6310..62231a6 100644 --- a/src/riscv/Platform/MetricSane.v +++ b/src/riscv/Platform/MetricSane.v @@ -4,6 +4,7 @@ Require Import coqutil.Tactics.Tactics. Require Import riscv.Spec.Machine. Require Import riscv.Utility.Monads. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Utility.Utility. Require Import riscv.Spec.Primitives. Require Import riscv.Platform.RiscvMachine. @@ -172,6 +173,7 @@ Section Sane. Proof. unfold run1. intros. apply Bind_sane; [apply getPC_sane|intros]. + apply Bind_sane; [apply leakEvent_sane|intros]. apply Bind_sane; [apply loadWord_sane|intros]. apply Bind_sane; [apply leakage_of_instr_sane|intros]. apply Bind_sane; [apply leakEvent_sane|intros]. diff --git a/src/riscv/Platform/Minimal.v b/src/riscv/Platform/Minimal.v index cdfbb44..9e7bd1a 100644 --- a/src/riscv/Platform/Minimal.v +++ b/src/riscv/Platform/Minimal.v @@ -5,6 +5,7 @@ Require Import Coq.Logic.PropExtensionality. Require Import riscv.Utility.Monads. Import OStateOperations. Require Import riscv.Utility.MonadNotations. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Spec.Machine. Require Import riscv.Utility.Utility. Require Import riscv.Spec.Primitives. @@ -45,7 +46,7 @@ Section Riscv. m <- fail_if_None (Memory.store_bytes n mach.(getMem) a v); update (fun mach => withXAddrs (invalidateWrittenXAddrs n a mach.(getXAddrs)) (withMem m mach)). - + Instance IsRiscvMachine: RiscvProgram (OState RiscvMachine) word := { getRegister reg := if Z.eq_dec reg Register0 then @@ -98,11 +99,16 @@ Section Riscv. (* fail hard if exception is thrown because at the moment, we want to prove that code output by the compiler never throws exceptions *) endCycleEarly{A: Type} := fail_hard; - }. - + }. Print leakEvent. Print withLeakageEvent. Print RiscvMachine. + Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (OState RiscvMachine) _ _ _ := { RVP := IsRiscvMachine; - leakEvent e := fail_hard; + leakEvent e := update (fun mach => + match e with + | Some e => withLeakageEvent e mach + (*I would write (anything mach) below, but universe constraints*) + | None => withLeakageEvent (anything (mach.(getRegs), mach.(getPc), mach.(getNextPc), mach.(getMem), mach.(getXAddrs), mach.(getLog))) mach + end); }. Arguments Memory.load_bytes: simpl never. @@ -230,6 +236,12 @@ Section Riscv. intros. eapply update_sane. intros. exists [e]. destruct mach. reflexivity. Qed. + Lemma leakEvent_sane: forall e, + mcomp_sane (leakEvent e). + Proof. + intros. eapply update_sane. intros. exists nil. destruct mach, e; reflexivity. + Qed. + Instance MinimalSane: PrimitivesSane MinimalPrimitivesParams. Proof. constructor. @@ -245,6 +257,7 @@ Section Riscv. all: repeat match goal with | |- _ => apply logEvent_sane + | |- _ => apply leakEvent_sane | |- mcomp_sane (Bind _ _) => apply Bind_sane | |- _ => apply Return_sane | |- _ => apply get_sane diff --git a/src/riscv/Platform/MinimalMMIO.v b/src/riscv/Platform/MinimalMMIO.v index 8f046df..84ed7cc 100644 --- a/src/riscv/Platform/MinimalMMIO.v +++ b/src/riscv/Platform/MinimalMMIO.v @@ -4,6 +4,7 @@ Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Require Export riscv.Utility.FreeMonad. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Spec.Machine. Require Import riscv.Utility.Utility. Require Import riscv.Spec.Primitives. @@ -114,7 +115,10 @@ Section Riscv. postF tt (withNextPc (word.add mach.(getPc) (word.of_Z 4)) mach) | EndCycleNormal => fun postF postA => postF tt (updatePc mach) | EndCycleEarly _ => fun postF postA => postA (updatePc mach) (* ignores postF containing the continuation *) - | LeakEvent e => fun postF postA => postF tt (withLeakageEvent e mach) + | LeakEvent e => fun postF postA => match e with + | Some e => postF tt (withLeakageEvent e mach) + | None => forall X (x : X), postF tt (withLeakageEvent (anything x) mach) + end | MakeReservation _ | ClearReservation _ | CheckReservation _ @@ -161,6 +165,7 @@ Section Riscv. Proof. destruct a; cbn; try solve [intuition eauto]. all : eauto using load_weaken_post, store_weaken_post. + intros. destruct o; eauto. Qed. Global Instance MinimalMMIOSatisfies_mcomp_sat_spec: mcomp_sat_spec MinimalMMIOPrimitivesParams. @@ -190,6 +195,7 @@ Section Riscv. cbv [load store nonmem_load nonmem_store]; cbn -[HList.tuple]; repeat destruct_one_match; intuition idtac; + try match goal with | H : forall _ _, _ |- _ => specialize (H True I) end; repeat lazymatch goal with | H : postF _ ?mach |- exists _ : RiscvMachine, _ => exists mach; cbn [RiscvMachine.getMem RiscvMachine.getXAddrs] @@ -203,7 +209,7 @@ Section Riscv. ssplit; eauto; simpl; change removeXAddr with (@List.removeb word word.eqb); rewrite ?ListSet.of_list_removeb; - intuition eauto 10 using preserve_undef_on, disjoint_diff_l. + intuition eauto 10 using preserve_undef_on, disjoint_diff_l. Qed. Lemma interpret_action_total'{memOk: map.ok Mem} a s post : From a757ed28fe91c870e538316f82f6836a36c08d4f Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 23:10:28 -0400 Subject: [PATCH 31/42] bring LeakageOfInstr (mostly) in compliance with Zkt spec. add comments. --- src/riscv/Spec/LeakageOfInstr.v | 234 ++++++++++++++------------------ 1 file changed, 102 insertions(+), 132 deletions(-) diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index 982fc22..c44c8e6 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -1,3 +1,24 @@ +(*To specify what is leaked by executing a RISC-V instruction, we attempt to follow + the RISC-V specification of the Zkt extension, as in The RISC-V Instruction Set Manual + Volume 1, Version 20240411. + + Ideally we'd like to say that if an implementation complies with Zkt, then instructions + leak no more than what we specify here. This is not quite true. + The one exception (so far), branching instructions, will be noted when it comes up. + + In all cases we specify, we assume that in the worst case, an instruction leaks itself and + all its arguments. That is, only register contents are leaked--not, for instance, memory + contents. + + Note that (as can be seen in Run.v), we always leak the full instruction; as the Zkt spec states, + "Binary executables should not contain secrets in the instruction encodings (Kerckhoffs’s + principle), so instruction timing may leak information about immediates, ordering of input + registers, etc." + Since the map (instruction represented as machine word -> Gallina representation of instruction) + defined in Decode.v is injective, it suffices to leak the Gallina representation, and this is + what we do. +*) + Require Import Utility.Utility. Require Import riscv.Utility.Monads. @@ -8,7 +29,7 @@ Notation Register := BinInt.Z (only parsing). Section WithMonad. Context {M : Type -> Type} {MM : Monad M}. - + Inductive LeakageM64 {width} {BW : Bitwidth width} {word: word.word width} : Type := | Mulw_leakage | Divw_leakage (num : word) (den : word) @@ -57,32 +78,32 @@ Section WithMonad. | Ld_leakage (addr: word) | Lwu_leakage (addr: word) | Addiw_leakage - | Slliw_leakage (shamt : Z) - | Srliw_leakage (shamt : Z) - | Sraiw_leakage (shamt : Z) + | Slliw_leakage + | Srliw_leakage + | Sraiw_leakage | Sd_leakage (addr: word) | Addw_leakage | Subw_leakage - | Sllw_leakage (shamt : word) - | Srlw_leakage (shamt : word) - | Sraw_leakage (shamt : word) + | Sllw_leakage + | Srlw_leakage + | Sraw_leakage | InvalidI64_leakage. Definition leakage_of_instr_I64 {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) (instr : InstructionI64) : M LeakageI64 := match instr with - | Ld _ rs1 oimm12 => addr <- getRegister rs1; Return (Ld_leakage (word.add addr (word.of_Z oimm12))) - | Lwu _ rs1 oimm12 => addr <- getRegister rs1; Return (Lwu_leakage (word.add addr (word.of_Z oimm12))) + | Ld _ rs1 _ => addr <- getRegister rs1; Return (Ld_leakage addr) + | Lwu _ rs1 _ => addr <- getRegister rs1; Return (Lwu_leakage addr) | Addiw _ _ _ => Return Addiw_leakage - | Slliw _ _ shamt => Return (Slliw_leakage shamt) - | Srliw _ _ shamt => Return (Srliw_leakage shamt) - | Sraiw _ _ shamt => Return (Sraiw_leakage shamt) - | Sd rs1 _ simm12 => addr <- getRegister rs1; Return (Sd_leakage (word.add addr (word.of_Z simm12))) + | Slliw _ _ _ => Return Slliw_leakage + | Srliw _ _ _ => Return Srliw_leakage + | Sraiw _ _ _ => Return Sraiw_leakage + | Sd rs1 _ _ => addr <- getRegister rs1; Return (Sd_leakage addr) | Addw _ _ _ => Return Addw_leakage | Subw _ _ _ => Return Subw_leakage - | Sllw _ _ rs2 => shamt <- getRegister rs2; Return (Sllw_leakage shamt) - | Srlw _ _ rs2 => shamt <- getRegister rs2; Return (Srlw_leakage shamt) - | Sraw _ _ rs2 => shamt <- getRegister rs2; Return (Sraw_leakage shamt) + | Sllw _ _ rs2 => Return Sllw_leakage + | Srlw _ _ rs2 => Return Srlw_leakage + | Sraw _ _ rs2 => Return Sraw_leakage | InvalidI64 => Return InvalidI64_leakage end. @@ -92,29 +113,29 @@ Section WithMonad. | Lw_leakage (addr: word) | Lbu_leakage (addr: word) | Lhu_leakage (addr: word) - | Fence_leakage (* unsure about this one. *) - | Fence_i_leakage + (*| Fence_leakage <- not implemented*) + (*| Fence_i_leakage <- not implemented*) | Addi_leakage - | Slli_leakage (shamt : Z) + | Slli_leakage | Slti_leakage | Sltiu_leakage | Xori_leakage | Ori_leakage | Andi_leakage - | Srli_leakage (shamt : Z) - | Srai_leakage (shamt : Z) + | Srli_leakage + | Srai_leakage | Auipc_leakage | Sb_leakage (addr: word) | Sh_leakage (addr: word) | Sw_leakage (addr: word) | Add_leakage | Sub_leakage - | Sll_leakage (shamt : word) + | Sll_leakage | Slt_leakage | Sltu_leakage | Xor_leakage - | Srl_leakage (shamt : word) - | Sra_leakage (shamt : word) + | Srl_leakage + | Sra_leakage | Or_leakage | And_leakage | Lui_leakage @@ -124,117 +145,62 @@ Section WithMonad. | Bge_leakage (branch: bool) | Bltu_leakage (branch: bool) | Bgeu_leakage (branch: bool) - | Jalr_leakage + | Jalr_leakage (addr : word) | Jal_leakage | InvalidI_leakage. + Notation "'ReturnSome' x" := (Return (Some x)) (at level 100). + + (*Here, we assume that branches leak only the value of the branch (i.e., yes or no) + and not the values being compared, although this is not stated in the spec.*) Definition leakage_of_instr_I {width} {BW : Bitwidth width} {word: word.word width} - (getRegister : Register -> M word) (instr : InstructionI) : M LeakageI := + (getRegister : Register -> M word) (instr : InstructionI) : M (option LeakageI) := match instr with - | Lb _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lb_leakage (word.add rs1_val (word.of_Z oimm12))) - | Lh _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lh_leakage (word.add rs1_val (word.of_Z oimm12))) - | Lw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lw_leakage (word.add rs1_val (word.of_Z oimm12))) - | Lbu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lbu_leakage (word.add rs1_val (word.of_Z oimm12))) - | Lhu _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Lhu_leakage (word.add rs1_val (word.of_Z oimm12))) - | Fence _ _ => Return Fence_leakage - | Fence_i => Return Fence_i_leakage - | Addi _ _ _ => Return Addi_leakage - | Slli _ _ shamt => Return (Slli_leakage shamt) - | Slti _ _ _ => Return Slti_leakage - | Sltiu _ _ _ => Return Sltiu_leakage - | Xori _ _ _ => Return Xori_leakage - | Ori _ _ _ => Return Ori_leakage - | Andi _ _ _ => Return Andi_leakage - | Srli _ _ shamt => Return (Srli_leakage shamt) - | Srai _ _ shamt => Return (Srai_leakage shamt) - | Auipc _ _ => Return Auipc_leakage - | Sb rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sb_leakage (word.add rs1_val (word.of_Z simm12))) - | Sh rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sh_leakage (word.add rs1_val (word.of_Z simm12))) - | Sw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Sw_leakage (word.add rs1_val (word.of_Z simm12))) - | Add _ _ _ => Return Add_leakage - | Sub _ _ _ => Return Sub_leakage - | Sll _ _ rs2 => shamt <- getRegister rs2; Return (Sll_leakage shamt) - | Slt _ _ _ => Return Slt_leakage - | Sltu _ _ _ => Return Sltu_leakage - | Xor _ _ _ => Return Xor_leakage - | Srl _ _ rs2 => shamt <- getRegister rs2; Return (Srl_leakage shamt) - | Sra _ _ rs2 => shamt <- getRegister rs2; Return (Sra_leakage shamt) - | Or _ _ _ => Return Or_leakage - | And _ _ _ => Return And_leakage - | Lui _ _ => Return Lui_leakage - | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Beq_leakage (word.eqb rs1_val rs2_val)) - | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bne_leakage (negb (word.eqb rs1_val rs2_val))) - | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Blt_leakage (word.lts rs1_val rs2_val)) - | Bge rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bge_leakage (negb (word.lts rs1_val rs2_val))) - | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bltu_leakage (word.ltu rs1_val rs2_val)) - | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; Return (Bgeu_leakage (negb (word.ltu rs1_val rs2_val))) - | Jalr _ _ _ => Return Jalr_leakage - | Jal _ _ => Return Jal_leakage - | InvalidI => Return InvalidI_leakage + | Lb _ rs1 _ => rs1_val <- getRegister rs1; ReturnSome (Lb_leakage rs1_val) + | Lh _ rs1 _ => rs1_val <- getRegister rs1; ReturnSome (Lh_leakage rs1_val) + | Lw _ rs1 _ => rs1_val <- getRegister rs1; ReturnSome (Lw_leakage rs1_val) + | Lbu _ rs1 _ => rs1_val <- getRegister rs1; ReturnSome(Lbu_leakage rs1_val) + | Lhu _ rs1 _ => rs1_val <- getRegister rs1; ReturnSome (Lhu_leakage rs1_val) + | Fence _ _ => Return None + | Fence_i => Return None + | Addi _ _ _ => ReturnSome Addi_leakage + | Slli _ _ _ => ReturnSome Slli_leakage + | Slti _ _ _ => ReturnSome Slti_leakage + | Sltiu _ _ _ => ReturnSome Sltiu_leakage + | Xori _ _ _ => ReturnSome Xori_leakage + | Ori _ _ _ => ReturnSome Ori_leakage + | Andi _ _ _ => ReturnSome Andi_leakage + | Srli _ _ _ => ReturnSome Srli_leakage + | Srai _ _ _ => ReturnSome Srai_leakage + | Auipc _ _ => ReturnSome Auipc_leakage + | Sb rs1 _ _ => rs1_val <- getRegister rs1; ReturnSome (Sb_leakage rs1_val) + | Sh rs1 _ _ => rs1_val <- getRegister rs1; ReturnSome (Sh_leakage rs1_val) + | Sw rs1 _ _ => rs1_val <- getRegister rs1; ReturnSome (Sw_leakage rs1_val) + | Add _ _ _ => ReturnSome Add_leakage + | Sub _ _ _ => ReturnSome Sub_leakage + | Sll _ _ rs2 => ReturnSome Sll_leakage + | Slt _ _ _ => ReturnSome Slt_leakage + | Sltu _ _ _ => ReturnSome Sltu_leakage + | Xor _ _ _ => ReturnSome Xor_leakage + | Srl _ _ rs2 => ReturnSome Srl_leakage + | Sra _ _ rs2 => ReturnSome Sra_leakage + | Or _ _ _ => ReturnSome Or_leakage + | And _ _ _ => ReturnSome And_leakage + | Lui _ _ => ReturnSome Lui_leakage + | Beq rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; ReturnSome (Beq_leakage (word.eqb rs1_val rs2_val)) + | Bne rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; ReturnSome (Bne_leakage (negb (word.eqb rs1_val rs2_val))) + | Blt rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; ReturnSome (Blt_leakage (word.lts rs1_val rs2_val)) + | Bge rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; ReturnSome (Bge_leakage (negb (word.lts rs1_val rs2_val))) + | Bltu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; ReturnSome (Bltu_leakage (word.ltu rs1_val rs2_val)) + | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; ReturnSome (Bgeu_leakage (negb (word.ltu rs1_val rs2_val))) + | Jalr _ rs1 _ => rs1_val <- getRegister rs1; ReturnSome (Jalr_leakage rs1_val) + | Jal _ _ => ReturnSome Jal_leakage + | InvalidI => ReturnSome InvalidI_leakage end. Inductive LeakageF64 : Type :=. - Inductive LeakageF {width} {BW : Bitwidth width} {word: word.word width} : Type := - | Flw_leakage (addr: word) - | Fsw_leakage (addr: word) - | Fmadd_s_leakage - | Fmsub_s_leakage - | Fnmsub_s_leakage - | Fnmadd_s_leakage - | Fadd_s_leakage - | Fsub_s_leakage - | Fmul_s_leakage - | Fdiv_s_leakage - | Fsqrt_s_leakage - | Fsgnj_s_leakage - | Fsgnjn_s_leakage - | Fsgnjx_s_leakage - | Fmin_s_leakage - | Fmax_s_leakage - | Fcvt_w_s_leakage - | Fcvt_wu_s_leakage - | Fmv_x_w_leakage - | Feq_s_leakage - | Flt_s_leakage - | Fle_s_leakage - | Fclass_s_leakage - | Fcvt_s_w_leakage - | Fcvt_s_wu_leakage - | Fmv_w_x_leakage - | InvalidF_leakage. - - Definition leakage_of_instr_F {width} {BW : Bitwidth width} {word: word.word width} - (getRegister : Register -> M word) (instr : InstructionF) : M LeakageF := - match instr with - | Flw _ rs1 oimm12 => rs1_val <- getRegister rs1; Return (Flw_leakage (word.add rs1_val (word.of_Z oimm12))) - | Fsw rs1 _ simm12 => rs1_val <- getRegister rs1; Return (Fsw_leakage (word.add rs1_val (word.of_Z simm12))) - | Fmadd_s _ _ _ _ _ => Return Fmadd_s_leakage - | Fmsub_s _ _ _ _ _ => Return Fmsub_s_leakage - | Fnmsub_s _ _ _ _ _ => Return Fnmsub_s_leakage - | Fnmadd_s _ _ _ _ _ => Return Fnmadd_s_leakage - | Fadd_s _ _ _ _ => Return Fadd_s_leakage - | Fsub_s _ _ _ _ => Return Fsub_s_leakage - | Fmul_s _ _ _ _ => Return Fmul_s_leakage - | Fdiv_s _ _ _ _ => Return Fdiv_s_leakage - | Fsqrt_s _ _ _ => Return Fsqrt_s_leakage - | Fsgnj_s _ _ _ => Return Fsgnj_s_leakage - | Fsgnjn_s _ _ _ => Return Fsgnjn_s_leakage - | Fsgnjx_s _ _ _ => Return Fsgnjx_s_leakage - | Fmin_s _ _ _ => Return Fmin_s_leakage - | Fmax_s _ _ _ => Return Fmax_s_leakage - | Fcvt_w_s _ _ _ => Return Fcvt_w_s_leakage - | Fcvt_wu_s _ _ _ => Return Fcvt_wu_s_leakage - | Fmv_x_w _ _ => Return Fmv_x_w_leakage - | Feq_s _ _ _ => Return Feq_s_leakage - | Flt_s _ _ _ => Return Flt_s_leakage - | Fle_s _ _ _ => Return Fle_s_leakage - | Fclass_s _ _ => Return Fclass_s_leakage - | Fcvt_s_w _ _ _ => Return Fcvt_s_w_leakage - | Fcvt_s_wu _ _ _ => Return Fcvt_s_wu_leakage - | Fmv_w_x _ _ => Return Fmv_w_x_leakage - | InvalidF => Return InvalidF_leakage - end. + Inductive LeakageF : Type :=. Inductive LeakageCSR : Type :=. @@ -259,16 +225,20 @@ Section WithMonad. Definition leakage_of_instr {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) (instr : Instruction) : M (option LeakageEvent) := match instr with - | IInstruction instr => l <- leakage_of_instr_I getRegister instr; Return (Some (ILeakage l)) - | MInstruction instr => l <- leakage_of_instr_M getRegister instr; Return (Some (MLeakage l)) + | IInstruction instr => l <- leakage_of_instr_I getRegister instr; + match l with + | Some l => ReturnSome (ILeakage l) + | None => Return None + end + | MInstruction instr => l <- leakage_of_instr_M getRegister instr; ReturnSome (MLeakage l) | AInstruction instr => Return None - | FInstruction instr => l <- leakage_of_instr_F getRegister instr; Return (Some (FLeakage l)) - | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; Return (Some (I64Leakage l)) - | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; Return (Some (M64Leakage l)) + | FInstruction instr => Return None + | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; ReturnSome (I64Leakage l) + | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; ReturnSome (M64Leakage l) | A64Instruction instr => Return None | F64Instruction instr => Return None | CSRInstruction instr => Return None - | InvalidInstruction _ => Return (Some InvalidLeakage) + | InvalidInstruction _ => ReturnSome InvalidLeakage end. End WithMonad. From 92a3acbd2b7f755e1fbb35dec9a324f796803f35 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 23:27:04 -0400 Subject: [PATCH 32/42] always leak the instruction --- src/riscv/Platform/Run.v | 4 ++-- src/riscv/Spec/LeakageOfInstr.v | 11 +++++++---- 2 files changed, 9 insertions(+), 6 deletions(-) diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index e777473..d71d602 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -25,8 +25,8 @@ Section Riscv. leakEvent (Some (fetchInstr pc));; inst <- loadWord Fetch pc; let inst' := decode iset (combine 4 inst) in - leakage_event <- leakage_of_instr getRegister inst'; - leakEvent leakage_event;; + inst_leakage <- leakage_of_instr getRegister inst'; + leakEvent (option_map (fun il => executeInstr inst' il) inst_leakage);; execute inst';; endCycleNormal. diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index c44c8e6..9be3198 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -208,9 +208,7 @@ Section WithMonad. Inductive LeakageA : Type :=. - Inductive LeakageEvent {width} {BW : Bitwidth width} {word: word.word width} : Type := - | anything {X : Type} (x : X) - | fetchInstr (address : word) + Inductive InstructionLeakage {width} {BW : Bitwidth width} {word: word.word width} : Type := | ILeakage (iLeakage : LeakageI) | MLeakage (mLeakage : LeakageM) | ALeakage (aLeakage : LeakageA) @@ -221,9 +219,14 @@ Section WithMonad. | F64Leakage (f64Leakage : LeakageF64) | CSRLeakage (csrLeakage : LeakageCSR) | InvalidLeakage. + + Inductive LeakageEvent {width} {BW : Bitwidth width} {word: word.word width} : Type := + | anything {X : Type} (x : X) + | fetchInstr (address : word) + | executeInstr (instr : Instruction) (ileakage : InstructionLeakage). Definition leakage_of_instr {width} {BW : Bitwidth width} {word: word.word width} - (getRegister : Register -> M word) (instr : Instruction) : M (option LeakageEvent) := + (getRegister : Register -> M word) (instr : Instruction) : M (option InstructionLeakage) := match instr with | IInstruction instr => l <- leakage_of_instr_I getRegister instr; match l with From 9b511c49882a14765fcc3346f26c582c8eaff6c4 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Sun, 4 Aug 2024 23:30:16 -0400 Subject: [PATCH 33/42] whitespace, etc --- src/riscv/Platform/Minimal.v | 4 ++-- src/riscv/Platform/MinimalMMIO.v | 2 +- src/riscv/Platform/MinimalMMIO_Post.v | 6 +++++- src/riscv/Platform/RiscvMachine.v | 6 +++--- src/riscv/Spec/Machine.v | 4 +--- 5 files changed, 12 insertions(+), 10 deletions(-) diff --git a/src/riscv/Platform/Minimal.v b/src/riscv/Platform/Minimal.v index 9e7bd1a..e2ee8c1 100644 --- a/src/riscv/Platform/Minimal.v +++ b/src/riscv/Platform/Minimal.v @@ -46,7 +46,7 @@ Section Riscv. m <- fail_if_None (Memory.store_bytes n mach.(getMem) a v); update (fun mach => withXAddrs (invalidateWrittenXAddrs n a mach.(getXAddrs)) (withMem m mach)). - + Instance IsRiscvMachine: RiscvProgram (OState RiscvMachine) word := { getRegister reg := if Z.eq_dec reg Register0 then @@ -99,7 +99,7 @@ Section Riscv. (* fail hard if exception is thrown because at the moment, we want to prove that code output by the compiler never throws exceptions *) endCycleEarly{A: Type} := fail_hard; - }. Print leakEvent. Print withLeakageEvent. Print RiscvMachine. + }. Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (OState RiscvMachine) _ _ _ := { RVP := IsRiscvMachine; diff --git a/src/riscv/Platform/MinimalMMIO.v b/src/riscv/Platform/MinimalMMIO.v index 84ed7cc..f93ff36 100644 --- a/src/riscv/Platform/MinimalMMIO.v +++ b/src/riscv/Platform/MinimalMMIO.v @@ -209,7 +209,7 @@ Section Riscv. ssplit; eauto; simpl; change removeXAddr with (@List.removeb word word.eqb); rewrite ?ListSet.of_list_removeb; - intuition eauto 10 using preserve_undef_on, disjoint_diff_l. + intuition eauto 10 using preserve_undef_on, disjoint_diff_l. Qed. Lemma interpret_action_total'{memOk: map.ok Mem} a s post : diff --git a/src/riscv/Platform/MinimalMMIO_Post.v b/src/riscv/Platform/MinimalMMIO_Post.v index f362135..d5459d6 100644 --- a/src/riscv/Platform/MinimalMMIO_Post.v +++ b/src/riscv/Platform/MinimalMMIO_Post.v @@ -3,6 +3,7 @@ Require Import Coq.ZArith.ZArith. Require Import riscv.Utility.Monads. Require Import riscv.Utility.MonadNotations. Require Import riscv.Spec.Decode. +Require Import riscv.Spec.LeakageOfInstr. Require Import riscv.Spec.Machine. Require Import riscv.Utility.Utility. Require Import riscv.Spec.Primitives. @@ -117,7 +118,10 @@ Section Riscv. Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (Post RiscvMachine) _ _ _ := {| RVP := IsRiscvMachine; - leakEvent _ := fun _ _ => False; + leakEvent e := fun mach post => match e with + | Some e => post tt (withLeakageEvent e mach) + | None => forall X (x : X), post tt (withLeakageEvent (anything x) mach) + end; |}. Definition MinimalMMIOPrimitivesParams: PrimitivesParams (Post RiscvMachine) RiscvMachine := {| diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index 5041e2f..08bc205 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -175,15 +175,15 @@ Section Machine. Definition withLogItems: list LogItem -> RiscvMachine -> RiscvMachine := fun items '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => - mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log) trace. + mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log) trace. Definition withLeakageEvent: LeakageEvent -> RiscvMachine -> RiscvMachine := fun event '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => - mkRiscvMachine regs pc nextPC mem xAddrs log (event :: trace). + mkRiscvMachine regs pc nextPC mem xAddrs log (event :: trace). Definition withLeakageEvents: list LeakageEvent -> RiscvMachine -> RiscvMachine := fun events '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => - mkRiscvMachine regs pc nextPC mem xAddrs log (events ++ trace). + mkRiscvMachine regs pc nextPC mem xAddrs log (events ++ trace). Definition Z32s_to_bytes(l: list Z): list byte := List.flat_map (fun z => HList.tuple.to_list (LittleEndian.split 4 z)) l. diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index c9fe29a..86cd837 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -67,9 +67,7 @@ Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { endCycleEarly: forall A, M A; }. -Class RiscvProgramWithLeakage - {width} {BW : Bitwidth width} {word: word.word width} - {M}{t}{MM}{MWt}:= +Class RiscvProgramWithLeakage{width}{BW : Bitwidth width}{word: word.word width}{M}{t}{MM}{MWt} := mkRiscvProgramWithLeakage { RVP :> @RiscvProgram M t MM MWt; leakEvent : (option LeakageEvent) -> M unit; From 1759c3b281dc73e6685c677fd8fe869ef5b5a433 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Mon, 5 Aug 2024 02:39:59 -0400 Subject: [PATCH 34/42] simplify Run.v --- src/riscv/Platform/Run.v | 4 ++-- src/riscv/Spec/LeakageOfInstr.v | 6 +++++- 2 files changed, 7 insertions(+), 3 deletions(-) diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index d71d602..e777473 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -25,8 +25,8 @@ Section Riscv. leakEvent (Some (fetchInstr pc));; inst <- loadWord Fetch pc; let inst' := decode iset (combine 4 inst) in - inst_leakage <- leakage_of_instr getRegister inst'; - leakEvent (option_map (fun il => executeInstr inst' il) inst_leakage);; + leakage_event <- leakage_of_instr getRegister inst'; + leakEvent leakage_event;; execute inst';; endCycleNormal. diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index 9be3198..8385b17 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -225,7 +225,7 @@ Section WithMonad. | fetchInstr (address : word) | executeInstr (instr : Instruction) (ileakage : InstructionLeakage). - Definition leakage_of_instr {width} {BW : Bitwidth width} {word: word.word width} + Definition instr_leakage {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) (instr : Instruction) : M (option InstructionLeakage) := match instr with | IInstruction instr => l <- leakage_of_instr_I getRegister instr; @@ -244,6 +244,10 @@ Section WithMonad. | InvalidInstruction _ => ReturnSome InvalidLeakage end. + Definition leakage_of_instr {width} {BW : Bitwidth width} {word: word.word width} + (getRegister : Register -> M word) (instr : Instruction) : M (option LeakageEvent) := + l <- instr_leakage getRegister instr; Return (option_map (executeInstr instr) l). + End WithMonad. Definition concrete_leakage_of_instr {width} {BW: Bitwidth width} {word: word.word width} := From a9305e0bbf665f29b23cf7bfe8e4fb7d9a535e57 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Mon, 5 Aug 2024 12:58:24 -0400 Subject: [PATCH 35/42] remove empty types --- src/riscv/Spec/LeakageOfInstr.v | 18 ++---------------- 1 file changed, 2 insertions(+), 16 deletions(-) diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index 8385b17..ae4c0e0 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -10,7 +10,8 @@ all its arguments. That is, only register contents are leaked--not, for instance, memory contents. - Note that (as can be seen in Run.v), we always leak the full instruction; as the Zkt spec states, + Note that (as can be seen in the definition of leakage_of_instr at the bottom of this file), + we always leak the full instruction; as the Zkt spec states, "Binary executables should not contain secrets in the instruction encodings (Kerckhoffs’s principle), so instruction timing may leak information about immediates, ordering of input registers, etc." @@ -198,26 +199,11 @@ Section WithMonad. | InvalidI => ReturnSome InvalidI_leakage end. - Inductive LeakageF64 : Type :=. - - Inductive LeakageF : Type :=. - - Inductive LeakageCSR : Type :=. - - Inductive LeakageA64 : Type :=. - - Inductive LeakageA : Type :=. - Inductive InstructionLeakage {width} {BW : Bitwidth width} {word: word.word width} : Type := | ILeakage (iLeakage : LeakageI) | MLeakage (mLeakage : LeakageM) - | ALeakage (aLeakage : LeakageA) - | FLeakage (fLeakage : LeakageF) | I64Leakage (i64Leakage : LeakageI64) | M64Leakage (m64Leakage : LeakageM64) - | A64Leakage (a64Leakage : LeakageA64) - | F64Leakage (f64Leakage : LeakageF64) - | CSRLeakage (csrLeakage : LeakageCSR) | InvalidLeakage. Inductive LeakageEvent {width} {BW : Bitwidth width} {word: word.word width} : Type := From 40b2ec82b1b52f48ae22945b19d1c4d27994ce97 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Mon, 5 Aug 2024 13:08:03 -0400 Subject: [PATCH 36/42] stop making claims about leakage of invalid instructions --- src/riscv/Spec/LeakageOfInstr.v | 91 ++++++++++++++++----------------- 1 file changed, 43 insertions(+), 48 deletions(-) diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index ae4c0e0..3f6471d 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -27,6 +27,7 @@ Require Import riscv.Utility.MonadNotations. Require Import riscv.Spec.Decode. Notation Register := BinInt.Z (only parsing). +Notation "'ReturnSome' x" := (Return (Some x)) (at level 100). Section WithMonad. Context {M : Type -> Type} {MM : Monad M}. @@ -37,17 +38,17 @@ Section WithMonad. | Divuw_leakage (num : word) (den : word) | Remw_leakage (num : word) (den : word) | Remuw_leakage (num : word) (den : word) - | InvalidM64_leakage. + (*| InvalidM64_leakage <- not specified*). Definition leakage_of_instr_M64 {width} {BW : Bitwidth width} {word: word.word width} - (getRegister : Register -> M word) (instr : InstructionM64) : M LeakageM64 := + (getRegister : Register -> M word) (instr : InstructionM64) : M (option LeakageM64) := match instr with - | Mulw _ _ _ => Return Mulw_leakage - | Divw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divw_leakage num den) - | Divuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divuw_leakage num den) - | Remw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remw_leakage num den) - | Remuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remuw_leakage num den) - | InvalidM64 => Return InvalidM64_leakage + | Mulw _ _ _ => ReturnSome Mulw_leakage + | Divw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; ReturnSome (Divw_leakage num den) + | Divuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; ReturnSome (Divuw_leakage num den) + | Remw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; ReturnSome (Remw_leakage num den) + | Remuw _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; ReturnSome (Remuw_leakage num den) + | InvalidM64 => Return None end. Inductive LeakageM {width} {BW : Bitwidth width} {word: word.word width} : Type := @@ -59,20 +60,20 @@ Section WithMonad. | Divu_leakage (num : word) (den : word) | Rem_leakage (num : word) (den : word) | Remu_leakage (num : word) (den : word) - | InvalidM_leakage. + (*| InvalidM_leakage*). Definition leakage_of_instr_M {width} {BW : Bitwidth width} {word: word.word width} - (getRegister : Register -> M word) (instr : InstructionM) : M LeakageM := + (getRegister : Register -> M word) (instr : InstructionM) : M (option LeakageM) := match instr with - | Mul _ _ _ => Return Mul_leakage - | Mulh _ _ _ => Return Mulh_leakage - | Mulhsu _ _ _ => Return Mulhsu_leakage - | Mulhu _ _ _ => Return Mulhu_leakage - | Div _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Div_leakage num den) - | Divu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Divu_leakage num den) - | Rem _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Rem_leakage num den) - | Remu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; Return (Remu_leakage num den) - | InvalidM => Return InvalidM_leakage + | Mul _ _ _ => ReturnSome Mul_leakage + | Mulh _ _ _ => ReturnSome Mulh_leakage + | Mulhsu _ _ _ => ReturnSome Mulhsu_leakage + | Mulhu _ _ _ => ReturnSome Mulhu_leakage + | Div _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; ReturnSome (Div_leakage num den) + | Divu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; ReturnSome (Divu_leakage num den) + | Rem _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; ReturnSome (Rem_leakage num den) + | Remu _ rs1 rs2 => num <- getRegister rs1; den <- getRegister rs2; ReturnSome (Remu_leakage num den) + | InvalidM => Return None end. Inductive LeakageI64 {width} {BW : Bitwidth width} {word: word.word width} : Type := @@ -88,24 +89,24 @@ Section WithMonad. | Sllw_leakage | Srlw_leakage | Sraw_leakage - | InvalidI64_leakage. + (*| InvalidI64_leakage*). Definition leakage_of_instr_I64 {width} {BW : Bitwidth width} {word: word.word width} - (getRegister : Register -> M word) (instr : InstructionI64) : M LeakageI64 := + (getRegister : Register -> M word) (instr : InstructionI64) : M (option LeakageI64) := match instr with - | Ld _ rs1 _ => addr <- getRegister rs1; Return (Ld_leakage addr) - | Lwu _ rs1 _ => addr <- getRegister rs1; Return (Lwu_leakage addr) - | Addiw _ _ _ => Return Addiw_leakage - | Slliw _ _ _ => Return Slliw_leakage - | Srliw _ _ _ => Return Srliw_leakage - | Sraiw _ _ _ => Return Sraiw_leakage - | Sd rs1 _ _ => addr <- getRegister rs1; Return (Sd_leakage addr) - | Addw _ _ _ => Return Addw_leakage - | Subw _ _ _ => Return Subw_leakage - | Sllw _ _ rs2 => Return Sllw_leakage - | Srlw _ _ rs2 => Return Srlw_leakage - | Sraw _ _ rs2 => Return Sraw_leakage - | InvalidI64 => Return InvalidI64_leakage + | Ld _ rs1 _ => addr <- getRegister rs1; ReturnSome (Ld_leakage addr) + | Lwu _ rs1 _ => addr <- getRegister rs1; ReturnSome (Lwu_leakage addr) + | Addiw _ _ _ => ReturnSome Addiw_leakage + | Slliw _ _ _ => ReturnSome Slliw_leakage + | Srliw _ _ _ => ReturnSome Srliw_leakage + | Sraiw _ _ _ => ReturnSome Sraiw_leakage + | Sd rs1 _ _ => addr <- getRegister rs1; ReturnSome (Sd_leakage addr) + | Addw _ _ _ => ReturnSome Addw_leakage + | Subw _ _ _ => ReturnSome Subw_leakage + | Sllw _ _ rs2 => ReturnSome Sllw_leakage + | Srlw _ _ rs2 => ReturnSome Srlw_leakage + | Sraw _ _ rs2 => ReturnSome Sraw_leakage + | InvalidI64 => Return None end. Inductive LeakageI {width} {BW : Bitwidth width} {word: word.word width} : Type := @@ -148,9 +149,7 @@ Section WithMonad. | Bgeu_leakage (branch: bool) | Jalr_leakage (addr : word) | Jal_leakage - | InvalidI_leakage. - - Notation "'ReturnSome' x" := (Return (Some x)) (at level 100). + (*| InvalidI_leakage*). (*Here, we assume that branches leak only the value of the branch (i.e., yes or no) and not the values being compared, although this is not stated in the spec.*) @@ -196,7 +195,7 @@ Section WithMonad. | Bgeu rs1 rs2 _ => rs1_val <- getRegister rs1; rs2_val <- getRegister rs2; ReturnSome (Bgeu_leakage (negb (word.ltu rs1_val rs2_val))) | Jalr _ rs1 _ => rs1_val <- getRegister rs1; ReturnSome (Jalr_leakage rs1_val) | Jal _ _ => ReturnSome Jal_leakage - | InvalidI => ReturnSome InvalidI_leakage + | InvalidI => Return None end. Inductive InstructionLeakage {width} {BW : Bitwidth width} {word: word.word width} : Type := @@ -204,7 +203,7 @@ Section WithMonad. | MLeakage (mLeakage : LeakageM) | I64Leakage (i64Leakage : LeakageI64) | M64Leakage (m64Leakage : LeakageM64) - | InvalidLeakage. + (*| InvalidLeakage*). Inductive LeakageEvent {width} {BW : Bitwidth width} {word: word.word width} : Type := | anything {X : Type} (x : X) @@ -214,20 +213,16 @@ Section WithMonad. Definition instr_leakage {width} {BW : Bitwidth width} {word: word.word width} (getRegister : Register -> M word) (instr : Instruction) : M (option InstructionLeakage) := match instr with - | IInstruction instr => l <- leakage_of_instr_I getRegister instr; - match l with - | Some l => ReturnSome (ILeakage l) - | None => Return None - end - | MInstruction instr => l <- leakage_of_instr_M getRegister instr; ReturnSome (MLeakage l) + | IInstruction instr => l <- leakage_of_instr_I getRegister instr; Return (option_map ILeakage l) + | MInstruction instr => l <- leakage_of_instr_M getRegister instr; Return (option_map MLeakage l) | AInstruction instr => Return None | FInstruction instr => Return None - | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; ReturnSome (I64Leakage l) - | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; ReturnSome (M64Leakage l) + | I64Instruction instr => l <- leakage_of_instr_I64 getRegister instr; Return (option_map I64Leakage l) + | M64Instruction instr => l <- leakage_of_instr_M64 getRegister instr; Return (option_map M64Leakage l) | A64Instruction instr => Return None | F64Instruction instr => Return None | CSRInstruction instr => Return None - | InvalidInstruction _ => ReturnSome InvalidLeakage + | InvalidInstruction _ => Return None end. Definition leakage_of_instr {width} {BW : Bitwidth width} {word: word.word width} From cc3c3fac70ce5ce97581f4f9e2aaa67b215caa19 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Mon, 5 Aug 2024 16:05:07 -0400 Subject: [PATCH 37/42] define LeakEvent in MinimalCSRs --- src/riscv/Platform/MinimalCSRs.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/riscv/Platform/MinimalCSRs.v b/src/riscv/Platform/MinimalCSRs.v index 1f61f08..f61a8f1 100644 --- a/src/riscv/Platform/MinimalCSRs.v +++ b/src/riscv/Platform/MinimalCSRs.v @@ -99,7 +99,7 @@ Section Riscv. | Machine => postF tt mach | User | Supervisor => False end - | LeakEvent _ + | LeakEvent _ => fun postF postA => postF tt mach | MakeReservation _ | ClearReservation _ | CheckReservation _ From 9949d74718347cf41bc13a6aba6eaf19a2b4ea1b Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Tue, 6 Aug 2024 00:07:33 -0400 Subject: [PATCH 38/42] simplify handling of instructions whose leakage trace has not been defined --- src/riscv/Platform/MetricRiscvMachine.v | 4 ++-- src/riscv/Platform/Minimal.v | 7 +------ src/riscv/Platform/MinimalMMIO.v | 7 +------ src/riscv/Platform/MinimalMMIO_Post.v | 5 +---- src/riscv/Platform/RiscvMachine.v | 20 ++++++++++++++------ src/riscv/Spec/LeakageOfInstr.v | 1 - src/riscv/Spec/MetricPrimitives.v | 5 +---- src/riscv/Spec/Primitives.v | 5 +---- 8 files changed, 21 insertions(+), 33 deletions(-) diff --git a/src/riscv/Platform/MetricRiscvMachine.v b/src/riscv/Platform/MetricRiscvMachine.v index 7a64963..77a28c6 100644 --- a/src/riscv/Platform/MetricRiscvMachine.v +++ b/src/riscv/Platform/MetricRiscvMachine.v @@ -59,11 +59,11 @@ Section Machine. fun items '(mkMetricRiscvMachine mach mc) => (mkMetricRiscvMachine (withLogItems items mach) mc). - Definition withLeakageEvent: LeakageEvent -> MetricRiscvMachine -> MetricRiscvMachine := + Definition withLeakageEvent: option LeakageEvent -> MetricRiscvMachine -> MetricRiscvMachine := fun event '(mkMetricRiscvMachine mach mc) => (mkMetricRiscvMachine (withLeakageEvent event mach) mc). - Definition withLeakageEvents: list LeakageEvent -> MetricRiscvMachine -> MetricRiscvMachine := + Definition withLeakageEvents: option (list LeakageEvent) -> MetricRiscvMachine -> MetricRiscvMachine := fun events '(mkMetricRiscvMachine mach mc) => (mkMetricRiscvMachine (withLeakageEvents events mach) mc). diff --git a/src/riscv/Platform/Minimal.v b/src/riscv/Platform/Minimal.v index e2ee8c1..8544233 100644 --- a/src/riscv/Platform/Minimal.v +++ b/src/riscv/Platform/Minimal.v @@ -103,12 +103,7 @@ Section Riscv. Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (OState RiscvMachine) _ _ _ := { RVP := IsRiscvMachine; - leakEvent e := update (fun mach => - match e with - | Some e => withLeakageEvent e mach - (*I would write (anything mach) below, but universe constraints*) - | None => withLeakageEvent (anything (mach.(getRegs), mach.(getPc), mach.(getNextPc), mach.(getMem), mach.(getXAddrs), mach.(getLog))) mach - end); + leakEvent e := update (withLeakageEvent e); }. Arguments Memory.load_bytes: simpl never. diff --git a/src/riscv/Platform/MinimalMMIO.v b/src/riscv/Platform/MinimalMMIO.v index f93ff36..f7a17dc 100644 --- a/src/riscv/Platform/MinimalMMIO.v +++ b/src/riscv/Platform/MinimalMMIO.v @@ -115,10 +115,7 @@ Section Riscv. postF tt (withNextPc (word.add mach.(getPc) (word.of_Z 4)) mach) | EndCycleNormal => fun postF postA => postF tt (updatePc mach) | EndCycleEarly _ => fun postF postA => postA (updatePc mach) (* ignores postF containing the continuation *) - | LeakEvent e => fun postF postA => match e with - | Some e => postF tt (withLeakageEvent e mach) - | None => forall X (x : X), postF tt (withLeakageEvent (anything x) mach) - end + | LeakEvent e => fun postF postA => postF tt (withLeakageEvent e mach) | MakeReservation _ | ClearReservation _ | CheckReservation _ @@ -165,7 +162,6 @@ Section Riscv. Proof. destruct a; cbn; try solve [intuition eauto]. all : eauto using load_weaken_post, store_weaken_post. - intros. destruct o; eauto. Qed. Global Instance MinimalMMIOSatisfies_mcomp_sat_spec: mcomp_sat_spec MinimalMMIOPrimitivesParams. @@ -195,7 +191,6 @@ Section Riscv. cbv [load store nonmem_load nonmem_store]; cbn -[HList.tuple]; repeat destruct_one_match; intuition idtac; - try match goal with | H : forall _ _, _ |- _ => specialize (H True I) end; repeat lazymatch goal with | H : postF _ ?mach |- exists _ : RiscvMachine, _ => exists mach; cbn [RiscvMachine.getMem RiscvMachine.getXAddrs] diff --git a/src/riscv/Platform/MinimalMMIO_Post.v b/src/riscv/Platform/MinimalMMIO_Post.v index d5459d6..0f7bc79 100644 --- a/src/riscv/Platform/MinimalMMIO_Post.v +++ b/src/riscv/Platform/MinimalMMIO_Post.v @@ -118,10 +118,7 @@ Section Riscv. Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (Post RiscvMachine) _ _ _ := {| RVP := IsRiscvMachine; - leakEvent e := fun mach post => match e with - | Some e => post tt (withLeakageEvent e mach) - | None => forall X (x : X), post tt (withLeakageEvent (anything x) mach) - end; + leakEvent e := fun mach post => post tt (withLeakageEvent e mach); |}. Definition MinimalMMIOPrimitivesParams: PrimitivesParams (Post RiscvMachine) RiscvMachine := {| diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index 08bc205..82d037a 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -142,7 +142,9 @@ Section Machine. getMem: Mem; getXAddrs: XAddrs; getLog: list LogItem; - getTrace: list LeakageEvent; + getTrace: option (list LeakageEvent); + (*^Some x means that only x has been leaked; + None means there are no guarantees at all about what has been leaked.*) }. Definition withRegs: Registers -> RiscvMachine -> RiscvMachine := @@ -165,6 +167,12 @@ Section Machine. fun xAddrs2 '(mkRiscvMachine regs pc nextPC mem xAddrs1 log trace) => mkRiscvMachine regs pc nextPC mem xAddrs2 log trace. + Definition option_map2 {X Y Z : Type} (f : X -> Y -> Z) x y := + match x, y with + | Some x, Some y => Some (f x y) + | _, _ => None + end. + Definition withLog: list LogItem -> RiscvMachine -> RiscvMachine := fun log2 '(mkRiscvMachine regs pc nextPC mem xAddrs log1 trace) => mkRiscvMachine regs pc nextPC mem xAddrs log2 trace. @@ -175,15 +183,15 @@ Section Machine. Definition withLogItems: list LogItem -> RiscvMachine -> RiscvMachine := fun items '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => - mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log) trace. + mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log) trace. - Definition withLeakageEvent: LeakageEvent -> RiscvMachine -> RiscvMachine := + Definition withLeakageEvent: option LeakageEvent -> RiscvMachine -> RiscvMachine := fun event '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => - mkRiscvMachine regs pc nextPC mem xAddrs log (event :: trace). + mkRiscvMachine regs pc nextPC mem xAddrs log (option_map2 cons event trace). - Definition withLeakageEvents: list LeakageEvent -> RiscvMachine -> RiscvMachine := + Definition withLeakageEvents: option (list LeakageEvent) -> RiscvMachine -> RiscvMachine := fun events '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => - mkRiscvMachine regs pc nextPC mem xAddrs log (events ++ trace). + mkRiscvMachine regs pc nextPC mem xAddrs log (option_map2 (@app _) events trace). Definition Z32s_to_bytes(l: list Z): list byte := List.flat_map (fun z => HList.tuple.to_list (LittleEndian.split 4 z)) l. diff --git a/src/riscv/Spec/LeakageOfInstr.v b/src/riscv/Spec/LeakageOfInstr.v index 3f6471d..7a6f5fa 100644 --- a/src/riscv/Spec/LeakageOfInstr.v +++ b/src/riscv/Spec/LeakageOfInstr.v @@ -206,7 +206,6 @@ Section WithMonad. (*| InvalidLeakage*). Inductive LeakageEvent {width} {BW : Bitwidth width} {word: word.word width} : Type := - | anything {X : Type} (x : X) | fetchInstr (address : word) | executeInstr (instr : Instruction) (ileakage : InstructionLeakage). diff --git a/src/riscv/Spec/MetricPrimitives.v b/src/riscv/Spec/MetricPrimitives.v index 927dad6..1fac880 100644 --- a/src/riscv/Spec/MetricPrimitives.v +++ b/src/riscv/Spec/MetricPrimitives.v @@ -138,10 +138,7 @@ Section MetricPrimitives. mcomp_sat endCycleNormal initialL post; spec_leakEvent: forall (initialL: MetricRiscvMachine) (post: unit -> MetricRiscvMachine -> Prop) (e : option LeakageEvent), - (match e with - | Some e => post tt (withLeakageEvent e initialL) - | None => forall X (x : X), post tt (withLeakageEvent (anything x) initialL) end) -> - mcomp_sat (leakEvent e) initialL post; + post tt (withLeakageEvent e initialL) -> mcomp_sat (leakEvent e) initialL post; }. End MetricPrimitives. diff --git a/src/riscv/Spec/Primitives.v b/src/riscv/Spec/Primitives.v index c6ea9e9..8bf58d8 100644 --- a/src/riscv/Spec/Primitives.v +++ b/src/riscv/Spec/Primitives.v @@ -178,10 +178,7 @@ Section Primitives. mcomp_sat endCycleNormal initialL post; spec_leakEvent: forall (initialL: RiscvMachine) (post: unit -> RiscvMachine -> Prop) (e : option LeakageEvent), - (match e with - | Some e => post tt (withLeakageEvent e initialL) - | None => forall X (x : X), post tt (withLeakageEvent (anything x) initialL) end) -> - mcomp_sat (leakEvent e) initialL post; + post tt (withLeakageEvent e initialL) -> mcomp_sat (leakEvent e) initialL post; }. End Primitives. From fc7e8a0a0201f2d1c4b833d29717919325b11130 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Tue, 6 Aug 2024 00:29:06 -0400 Subject: [PATCH 39/42] whitespace, etc --- src/riscv/Platform/Minimal.v | 2 +- src/riscv/Platform/RiscvMachine.v | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/riscv/Platform/Minimal.v b/src/riscv/Platform/Minimal.v index 8544233..11104f2 100644 --- a/src/riscv/Platform/Minimal.v +++ b/src/riscv/Platform/Minimal.v @@ -234,7 +234,7 @@ Section Riscv. Lemma leakEvent_sane: forall e, mcomp_sane (leakEvent e). Proof. - intros. eapply update_sane. intros. exists nil. destruct mach, e; reflexivity. + intros. eapply update_sane. intros. exists nil. destruct mach. reflexivity. Qed. Instance MinimalSane: PrimitivesSane MinimalPrimitivesParams. diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index 82d037a..83f314f 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -183,7 +183,7 @@ Section Machine. Definition withLogItems: list LogItem -> RiscvMachine -> RiscvMachine := fun items '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => - mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log) trace. + mkRiscvMachine regs pc nextPC mem xAddrs (items ++ log) trace. Definition withLeakageEvent: option LeakageEvent -> RiscvMachine -> RiscvMachine := fun event '(mkRiscvMachine regs pc nextPC mem xAddrs log trace) => From 651109955686b3b64374bded2d65cc13a39d6705 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Tue, 6 Aug 2024 13:40:38 -0400 Subject: [PATCH 40/42] make type inference nicer with coercion and different implicits. quiet warnings: replace ":>" with "::" and "revert dependent" with "generalize dependent" --- src/riscv/Platform/MaterializeRiscvProgram.v | 2 +- .../Platform/MetricMaterializeRiscvProgram.v | 2 +- src/riscv/Platform/MetricMinimal.v | 2 +- src/riscv/Platform/MetricSane.v | 2 +- src/riscv/Platform/Minimal.v | 2 +- src/riscv/Platform/MinimalMMIO_Post.v | 2 +- src/riscv/Platform/Run.v | 2 +- src/riscv/Spec/Machine.v | 14 ++++++++----- src/riscv/Spec/MetricPrimitives.v | 20 +++++++++---------- src/riscv/Spec/Primitives.v | 20 +++++++++---------- src/riscv/Utility/FreeMonad.v | 4 ++-- 11 files changed, 38 insertions(+), 34 deletions(-) diff --git a/src/riscv/Platform/MaterializeRiscvProgram.v b/src/riscv/Platform/MaterializeRiscvProgram.v index f71add6..8270d6f 100644 --- a/src/riscv/Platform/MaterializeRiscvProgram.v +++ b/src/riscv/Platform/MaterializeRiscvProgram.v @@ -87,7 +87,7 @@ Section Riscv. endCycleEarly A := act (EndCycleEarly A) ret; |}. - Global Instance MaterializeWithLeakage : RiscvProgramWithLeakage := {| + Global Instance MaterializeWithLeakage : RiscvProgramWithLeakage (free riscv_primitive primitive_result) word := {| RVP := Materialize; leakEvent a := act (LeakEvent a) ret |}. diff --git a/src/riscv/Platform/MetricMaterializeRiscvProgram.v b/src/riscv/Platform/MetricMaterializeRiscvProgram.v index 4335be2..23a096d 100644 --- a/src/riscv/Platform/MetricMaterializeRiscvProgram.v +++ b/src/riscv/Platform/MetricMaterializeRiscvProgram.v @@ -39,7 +39,7 @@ Section Riscv. endCycleEarly A := act (addMetricInstructions 1, EndCycleEarly A) ret; |}. - Global Instance MetricMaterializeWithLeakage : RiscvProgramWithLeakage := {| + Global Instance MetricMaterializeWithLeakage : RiscvProgramWithLeakage (free action result) word := {| RVP := MetricMaterialize; leakEvent a := act (id, LeakEvent a) ret |}. diff --git a/src/riscv/Platform/MetricMinimal.v b/src/riscv/Platform/MetricMinimal.v index f6e1c60..570388b 100644 --- a/src/riscv/Platform/MetricMinimal.v +++ b/src/riscv/Platform/MetricMinimal.v @@ -66,7 +66,7 @@ Section Riscv. endCycleEarly{A} := liftL0 (addMetricInstructions 1) (@endCycleEarly _ _ _ _ _ A); }. - Instance IsMetricRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (OState MetricRiscvMachine) _ _ _ := { + Instance IsMetricRiscvMachineWithLeakage: RiscvProgramWithLeakage (OState MetricRiscvMachine) word := { RVP := IsMetricRiscvMachine; leakEvent := liftL1 id leakEvent; }. diff --git a/src/riscv/Platform/MetricSane.v b/src/riscv/Platform/MetricSane.v index 62231a6..04d28ae 100644 --- a/src/riscv/Platform/MetricSane.v +++ b/src/riscv/Platform/MetricSane.v @@ -21,7 +21,7 @@ Section Sane. Context {mem: map.map word byte}. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgramWithLeakage}. + Context {RVM: RiscvProgramWithLeakage M word}. Context {PRParams: PrimitivesParams M MetricRiscvMachine}. Context {mcomp_sat_ok: mcomp_sat_spec PRParams}. diff --git a/src/riscv/Platform/Minimal.v b/src/riscv/Platform/Minimal.v index 11104f2..f30980f 100644 --- a/src/riscv/Platform/Minimal.v +++ b/src/riscv/Platform/Minimal.v @@ -101,7 +101,7 @@ Section Riscv. endCycleEarly{A: Type} := fail_hard; }. - Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (OState RiscvMachine) _ _ _ := { + Instance IsRiscvMachineWithLeakage: RiscvProgramWithLeakage (OState RiscvMachine) word := { RVP := IsRiscvMachine; leakEvent e := update (withLeakageEvent e); }. diff --git a/src/riscv/Platform/MinimalMMIO_Post.v b/src/riscv/Platform/MinimalMMIO_Post.v index 0f7bc79..f03f04b 100644 --- a/src/riscv/Platform/MinimalMMIO_Post.v +++ b/src/riscv/Platform/MinimalMMIO_Post.v @@ -116,7 +116,7 @@ Section Riscv. - exact (post tt (withPc mach.(getNextPc) (withNextPc (word.add mach.(getNextPc) (word.of_Z 4)) mach))). Defined. - Instance IsRiscvMachineWithLeakage: @RiscvProgramWithLeakage _ _ _ (Post RiscvMachine) _ _ _ := {| + Instance IsRiscvMachineWithLeakage: RiscvProgramWithLeakage (Post RiscvMachine) word := {| RVP := IsRiscvMachine; leakEvent e := fun mach post => post tt (withLeakageEvent e mach); |}. diff --git a/src/riscv/Platform/Run.v b/src/riscv/Platform/Run.v index e777473..48f277e 100644 --- a/src/riscv/Platform/Run.v +++ b/src/riscv/Platform/Run.v @@ -16,7 +16,7 @@ Section Riscv. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgramWithLeakage}. + Context {RVM: RiscvProgramWithLeakage M mword}. Context {RVS: RiscvMachine M mword}. Definition run1(iset: InstructionSet): diff --git a/src/riscv/Spec/Machine.v b/src/riscv/Spec/Machine.v index 86cd837..65faddd 100644 --- a/src/riscv/Spec/Machine.v +++ b/src/riscv/Spec/Machine.v @@ -67,11 +67,13 @@ Class RiscvProgram{M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgram { endCycleEarly: forall A, M A; }. -Class RiscvProgramWithLeakage{width}{BW : Bitwidth width}{word: word.word width}{M}{t}{MM}{MWt} := - mkRiscvProgramWithLeakage { - RVP :> @RiscvProgram M t MM MWt; - leakEvent : (option LeakageEvent) -> M unit; - }. +Class RiscvProgramWithLeakage{width}{BW : Bitwidth width}{word: word.word width} + {M}{t}`{Monad M}`{MachineWidth t} := mkRiscvProgramWithLeakage { + RVP :: RiscvProgram; + leakEvent : (option LeakageEvent) -> M unit; +}. + +Coercion RVP : RiscvProgramWithLeakage >-> RiscvProgram. Class RiscvMachine`{MP: RiscvProgram} := mkRiscvMachine { (* checks that addr is aligned, and translates the (possibly virtual) addr to a physical @@ -150,6 +152,8 @@ Notation Register0 := 0%Z (only parsing). Arguments RiscvProgram: clear implicits. Arguments RiscvProgram (M) (t) {_} {_}. +Arguments RiscvProgramWithLeakage: clear implicits. +Arguments RiscvProgramWithLeakage {_} {_} {_} (M) (t) {_} {_}. Arguments RiscvMachine: clear implicits. Arguments RiscvMachine (M) (t) {_} {_} {_}. diff --git a/src/riscv/Spec/MetricPrimitives.v b/src/riscv/Spec/MetricPrimitives.v index 1fac880..933304d 100644 --- a/src/riscv/Spec/MetricPrimitives.v +++ b/src/riscv/Spec/MetricPrimitives.v @@ -21,8 +21,8 @@ Section MetricPrimitives. Context {M: Type -> Type}. Context {MM: Monad M}. - Context {RVM: RiscvProgramWithLeakage}. - Context {RVS: @RiscvMachine M word _ _ RVM.(RVP)}. + Context {RVM: RiscvProgramWithLeakage M word}. + Context {RVS: @RiscvMachine M word _ _ RVM}. (* monadic computations used for specifying the behavior of RiscvMachines should be "sane" in the sense that we never step to the empty set (that's not absence of failure, since @@ -110,15 +110,15 @@ Section MetricPrimitives. x = Register0 /\ post tt initialL) -> mcomp_sat (setRegister x v) initialL post; - spec_loadByte: spec_load 1 (Machine.loadByte (RiscvProgram := RVM.(RVP))) Memory.loadByte; - spec_loadHalf: spec_load 2 (Machine.loadHalf (RiscvProgram := RVM.(RVP))) Memory.loadHalf; - spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM.(RVP))) Memory.loadWord; - spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM.(RVP))) Memory.loadDouble; + spec_loadByte: spec_load 1 (Machine.loadByte (RiscvProgram := RVM)) Memory.loadByte; + spec_loadHalf: spec_load 2 (Machine.loadHalf (RiscvProgram := RVM)) Memory.loadHalf; + spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM)) Memory.loadWord; + spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM)) Memory.loadDouble; - spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM.(RVP))) Memory.storeByte; - spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM.(RVP))) Memory.storeHalf; - spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM.(RVP))) Memory.storeWord; - spec_storeDouble: spec_store 8 (Machine.storeDouble (RiscvProgram := RVM.(RVP))) Memory.storeDouble; + spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM)) Memory.storeByte; + spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM)) Memory.storeHalf; + spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM)) Memory.storeWord; + spec_storeDouble: spec_store 8 (Machine.storeDouble (RiscvProgram := RVM)) Memory.storeDouble; spec_getPC: forall (initialL: MetricRiscvMachine) (post: word -> MetricRiscvMachine -> Prop), post initialL.(getPc) initialL -> diff --git a/src/riscv/Spec/Primitives.v b/src/riscv/Spec/Primitives.v index 8bf58d8..e2fc3ba 100644 --- a/src/riscv/Spec/Primitives.v +++ b/src/riscv/Spec/Primitives.v @@ -71,8 +71,8 @@ Section Primitives. (mcomp_sat comp st (fun a st' => (post a st' /\ exists diff, st'.(getLog) = diff ++ st.(getLog)) /\ valid_machine st')). - Context {RVM: RiscvProgramWithLeakage}. - Context {RVS: @riscv.Spec.Machine.RiscvMachine M word _ _ RVM.(RVP)}. + Context {RVM: RiscvProgramWithLeakage M word}. + Context {RVS: @riscv.Spec.Machine.RiscvMachine M word _ _ RVM}. Class PrimitivesSane(p: PrimitivesParams RiscvMachine): Prop := { getRegister_sane: forall r, mcomp_sane (getRegister r); @@ -153,15 +153,15 @@ Section Primitives. x = Register0 /\ post tt initialL) -> mcomp_sat (setRegister x v) initialL post; - spec_loadByte: spec_load 1 (Machine.loadByte (RiscvProgram := RVM.(RVP))) Memory.loadByte; - spec_loadHalf: spec_load 2 (Machine.loadHalf (RiscvProgram := RVM.(RVP))) Memory.loadHalf; - spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM.(RVP))) Memory.loadWord; - spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM.(RVP))) Memory.loadDouble; + spec_loadByte: spec_load 1 (Machine.loadByte (RiscvProgram := RVM)) Memory.loadByte; + spec_loadHalf: spec_load 2 (Machine.loadHalf (RiscvProgram := RVM)) Memory.loadHalf; + spec_loadWord: spec_load 4 (Machine.loadWord (RiscvProgram := RVM)) Memory.loadWord; + spec_loadDouble: spec_load 8 (Machine.loadDouble (RiscvProgram := RVM)) Memory.loadDouble; - spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM.(RVP))) Memory.storeByte; - spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM.(RVP))) Memory.storeHalf; - spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM.(RVP))) Memory.storeWord; - spec_storeDouble: spec_store 8 (Machine.storeDouble (RiscvProgram := RVM.(RVP))) Memory.storeDouble; + spec_storeByte: spec_store 1 (Machine.storeByte (RiscvProgram := RVM)) Memory.storeByte; + spec_storeHalf: spec_store 2 (Machine.storeHalf (RiscvProgram := RVM)) Memory.storeHalf; + spec_storeWord: spec_store 4 (Machine.storeWord (RiscvProgram := RVM)) Memory.storeWord; + spec_storeDouble: spec_store 8 (Machine.storeDouble (RiscvProgram := RVM)) Memory.storeDouble; spec_getPC: forall initialL (post: word -> RiscvMachine -> Prop), post initialL.(getPc) initialL -> diff --git a/src/riscv/Utility/FreeMonad.v b/src/riscv/Utility/FreeMonad.v index 5dfb335..f74bea7 100644 --- a/src/riscv/Utility/FreeMonad.v +++ b/src/riscv/Utility/FreeMonad.v @@ -56,7 +56,7 @@ Module free. Lemma interp_weaken_post {T} (p : free T) s (post1 post2:_->_->Prop) (Hpost : forall r s, post1 r s -> post2 r s) (Hinterp : interp p s post1) : interp p s post2. - Proof. revert dependent s; induction p; cbn; firstorder eauto. Qed. + Proof. generalize dependent s; induction p; cbn; firstorder eauto. Qed. Lemma interp_bind {A B} s post (a : free A) (b : A -> free B) : interp (bind a b) s post <-> interp a s (fun x s => interp (b x) s post). @@ -115,7 +115,7 @@ Module free. (forall s, postA1 s -> postA2 s) -> interpret p s postF1 postA1 -> interpret p s postF2 postA2. Proof. - revert dependent s; induction p. 2: firstorder eauto. + generalize dependent s; induction p. 2: firstorder eauto. cbn. intros. eapply interpret_action_weaken_post. - intros. eapply H; try eassumption. exact H3. - eassumption. From 1d4b2d15ed4f0314cbe875f3e2a497818401c276 Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Wed, 7 Aug 2024 14:30:16 -0400 Subject: [PATCH 41/42] import option_map2 instead of defining it --- src/riscv/Platform/RiscvMachine.v | 7 +------ 1 file changed, 1 insertion(+), 6 deletions(-) diff --git a/src/riscv/Platform/RiscvMachine.v b/src/riscv/Platform/RiscvMachine.v index 83f314f..6bb4a4b 100644 --- a/src/riscv/Platform/RiscvMachine.v +++ b/src/riscv/Platform/RiscvMachine.v @@ -1,5 +1,6 @@ Require Import Coq.Strings.String. Require Import Coq.ZArith.BinInt. +Require Import coqutil.Datatypes.Option. Require Import coqutil.Map.Interface. Require Import coqutil.Word.Interface. Require Import coqutil.Word.LittleEndian. @@ -167,12 +168,6 @@ Section Machine. fun xAddrs2 '(mkRiscvMachine regs pc nextPC mem xAddrs1 log trace) => mkRiscvMachine regs pc nextPC mem xAddrs2 log trace. - Definition option_map2 {X Y Z : Type} (f : X -> Y -> Z) x y := - match x, y with - | Some x, Some y => Some (f x y) - | _, _ => None - end. - Definition withLog: list LogItem -> RiscvMachine -> RiscvMachine := fun log2 '(mkRiscvMachine regs pc nextPC mem xAddrs log1 trace) => mkRiscvMachine regs pc nextPC mem xAddrs log2 trace. From 8587bd40439f9f8c995b206b78b4100024bbde4f Mon Sep 17 00:00:00 2001 From: Owen Conoly Date: Wed, 7 Aug 2024 14:42:50 -0400 Subject: [PATCH 42/42] fix Examples --- src/riscv/Examples/Example64Literal.v | 1 + src/riscv/Examples/Fib.v | 1 + 2 files changed, 2 insertions(+) diff --git a/src/riscv/Examples/Example64Literal.v b/src/riscv/Examples/Example64Literal.v index 2670930..4be5186 100644 --- a/src/riscv/Examples/Example64Literal.v +++ b/src/riscv/Examples/Example64Literal.v @@ -46,6 +46,7 @@ Definition zeroedRiscvMachine: RiscvMachine := {| getMem := map.empty; getXAddrs := nil; getLog := nil; + getTrace := Some nil; |}. Definition initialRiscvMachine(imem: list MachineInt): RiscvMachine := diff --git a/src/riscv/Examples/Fib.v b/src/riscv/Examples/Fib.v index 68381f1..2074c24 100644 --- a/src/riscv/Examples/Fib.v +++ b/src/riscv/Examples/Fib.v @@ -57,6 +57,7 @@ Definition zeroedRiscvMachine: RiscvMachine := {| getMem := map.empty; getXAddrs := nil; getLog := nil; + getTrace := Some nil; |}. Definition initialRiscvMachine(imem: list MachineInt): RiscvMachine :=