feat(Frontend): support hardware error exception

This commit is contained in:
ngc7331 2025-06-03 14:45:30 +08:00
parent 6b4ab11c1e
commit ec2ac1dfc9
No known key found for this signature in database
GPG Key ID: FE4095A6F75802D0
9 changed files with 169 additions and 141 deletions

View File

@ -174,6 +174,8 @@ class ExceptionType extends Bundle {
def isPf: Bool = value === ExceptionType.Value.Pf def isPf: Bool = value === ExceptionType.Value.Pf
def isGpf: Bool = value === ExceptionType.Value.Gpf def isGpf: Bool = value === ExceptionType.Value.Gpf
def isAf: Bool = value === ExceptionType.Value.Af def isAf: Bool = value === ExceptionType.Value.Af
def isIll: Bool = value === ExceptionType.Value.Ill
def isHwe: Bool = value === ExceptionType.Value.Hwe
def hasException: Bool = value =/= ExceptionType.Value.None def hasException: Bool = value =/= ExceptionType.Value.None
@ -191,11 +193,13 @@ class ExceptionType extends Bundle {
} }
object ExceptionType { object ExceptionType {
private object Value extends EnumUInt(4) { private object Value extends EnumUInt(6) {
def None: UInt = 0.U(width.W) def None: UInt = 0.U(width.W)
def Pf: UInt = 1.U(width.W) // instruction page fault def Pf: UInt = 1.U(width.W) // instruction page fault
def Gpf: UInt = 2.U(width.W) // instruction guest page fault def Gpf: UInt = 2.U(width.W) // instruction guest page fault
def Af: UInt = 3.U(width.W) // instruction access fault def Af: UInt = 3.U(width.W) // instruction access fault
def Ill: UInt = 4.U(width.W) // illegal instruction
def Hwe: UInt = 5.U(width.W) // hardware error
} }
def apply(that: UInt): ExceptionType = { def apply(that: UInt): ExceptionType = {
@ -209,14 +213,25 @@ object ExceptionType {
def Pf: ExceptionType = apply(Value.Pf) def Pf: ExceptionType = apply(Value.Pf)
def Gpf: ExceptionType = apply(Value.Gpf) def Gpf: ExceptionType = apply(Value.Gpf)
def Af: ExceptionType = apply(Value.Af) def Af: ExceptionType = apply(Value.Af)
def Ill: ExceptionType = apply(Value.Ill)
def Hwe: ExceptionType = apply(Value.Hwe)
def apply(hasPf: Bool, hasGpf: Bool, hasAf: Bool): ExceptionType = { def apply(
hasPf: Bool,
hasGpf: Bool,
hasAf: Bool,
hasIll: Bool,
hasHwe: Bool,
canAssert: Bool = true.B // prevent assert(x)
): ExceptionType = {
assert( assert(
PopCount(VecInit(hasPf, hasGpf, hasAf)) <= 1.U, !canAssert || PopCount(VecInit(hasPf, hasGpf, hasAf, hasIll, hasHwe)) <= 1.U,
"ExceptionType receives input that is not one-hot: pf=%d, gpf=%d, af=%d", "ExceptionType receives input that is not one-hot: pf=%d, gpf=%d, af=%d, ill=%d, hwe=%d",
hasPf, hasPf,
hasGpf, hasGpf,
hasAf hasAf,
hasIll,
hasHwe
) )
// input is at-most-one-hot encoded, so we don't worry about priority here. // input is at-most-one-hot encoded, so we don't worry about priority here.
MuxCase( MuxCase(
@ -224,38 +239,60 @@ object ExceptionType {
Seq( Seq(
hasPf -> Pf, hasPf -> Pf,
hasGpf -> Gpf, hasGpf -> Gpf,
hasAf -> Af hasAf -> Af,
hasIll -> Ill,
hasHwe -> Hwe
) )
) )
} }
// only af is used most frequently (pmp / ecc / tilelink), so we define a shortcut // only af is used most frequently (pmp / ecc / tilelink), so we define a shortcut
// we cannot use default parameter in apply(), as it is overloaded and scala won't allow // we cannot use default parameter in apply(), as it is overloaded and scala won't allow
def apply(hasAf: Bool): ExceptionType = def hasAf(hasAf: Bool, canAssert: Bool = true.B): ExceptionType =
apply(hasPf = false.B, hasGpf = false.B, hasAf = hasAf) apply(false.B, false.B, hasAf, false.B, false.B, canAssert)
// raise pf/gpf/af according to backend redirect request (tlb pre-check) // raise pf/gpf/af according to backend redirect request (tlb pre-check)
def fromBackend(redirect: Redirect): ExceptionType = def fromBackend(redirect: Redirect, canAssert: Bool = true.B): ExceptionType =
apply( apply(
hasPf = redirect.backendIPF, redirect.backendIPF,
hasGpf = redirect.backendIGPF, redirect.backendIGPF,
hasAf = redirect.backendIAF redirect.backendIAF,
false.B,
false.B,
canAssert
) )
// raise pf/gpf/af according to itlb response // raise pf/gpf/af according to itlb response
def fromTlbResp(resp: TlbResp, useDup: Int = 0): ExceptionType = { def fromTlbResp(resp: TlbResp, canAssert: Bool = true.B, useDup: Int = 0): ExceptionType = {
require(useDup >= 0 && useDup < resp.excp.length) require(useDup >= 0 && useDup < resp.excp.length)
// itlb is guaranteed to respond at most one exception // itlb is guaranteed to respond at most one exception
apply( apply(
hasPf = resp.excp(useDup).pf.instr, resp.excp(useDup).pf.instr,
hasGpf = resp.excp(useDup).gpf.instr, resp.excp(useDup).gpf.instr,
hasAf = resp.excp(useDup).af.instr resp.excp(useDup).af.instr,
false.B,
false.B,
canAssert
) )
} }
// raise af if pmp check failed // raise af if pmp check failed
def fromPmpResp(resp: PMPRespBundle): ExceptionType = def fromPmpResp(resp: PMPRespBundle, canAssert: Bool = true.B): ExceptionType =
apply(hasAf = resp.instr) hasAf(resp.instr, canAssert)
// raise af / hwe according to tilelink response
def fromTileLink(corrupt: Bool, denied: Bool, canAssert: Bool = true.B): ExceptionType = {
// corrupt && denied -> access fault
// corrupt && !denied -> hardware error
// !corrupt && !denied -> no exception
// !corrupt && denied -> violates tilelink specification, should not happen
assert(!canAssert || !(denied && !corrupt), "TileLink response should not be denied but !corrupt")
apply(false.B, false.B, denied, false.B, corrupt && !denied, canAssert)
}
// raise ill according to rvc expander
def fromRvcExpander(ill: Bool, canAssert: Bool = true.B): ExceptionType =
apply(false.B, false.B, false.B, ill, false.B, canAssert)
} }
object BrType extends EnumUInt(4) { object BrType extends EnumUInt(4) {
@ -297,12 +334,11 @@ class FetchToIBuffer(implicit p: Parameters) extends FrontendBundle {
val foldpc: Vec[UInt] = Vec(IBufferEnqueueWidth, UInt(MemPredPCWidth.W)) val foldpc: Vec[UInt] = Vec(IBufferEnqueueWidth, UInt(MemPredPCWidth.W))
val instrEndOffset: Vec[InstrEndOffset] = Vec(IBufferEnqueueWidth, new InstrEndOffset) val instrEndOffset: Vec[InstrEndOffset] = Vec(IBufferEnqueueWidth, new InstrEndOffset)
// val ftqPcOffset: Vec[Valid[FtqPcOffset]] = Vec(IBufferEnqueueWidth, Valid(new FtqPcOffset)) // val ftqPcOffset: Vec[Valid[FtqPcOffset]] = Vec(IBufferEnqueueWidth, Valid(new FtqPcOffset))
val backendException: Vec[Bool] = Vec(IBufferEnqueueWidth, Bool()) val isBackendException: Vec[Bool] = Vec(IBufferEnqueueWidth, Bool())
val exceptionType: Vec[ExceptionType] = Vec(IBufferEnqueueWidth, new ExceptionType) val exceptionType: Vec[ExceptionType] = Vec(IBufferEnqueueWidth, new ExceptionType)
val crossPageIPFFix: Vec[Bool] = Vec(IBufferEnqueueWidth, Bool()) val exceptionCrossPage: Vec[Bool] = Vec(IBufferEnqueueWidth, Bool())
val illegalInstr: Vec[Bool] = Vec(IBufferEnqueueWidth, Bool()) val triggered: Vec[UInt] = Vec(IBufferEnqueueWidth, TriggerAction())
val triggered: Vec[UInt] = Vec(IBufferEnqueueWidth, TriggerAction()) val isLastInFtqEntry: Vec[Bool] = Vec(IBufferEnqueueWidth, Bool())
val isLastInFtqEntry: Vec[Bool] = Vec(IBufferEnqueueWidth, Bool())
val pc: Vec[PrunedAddr] = Vec(IBufferEnqueueWidth, PrunedAddr(VAddrBits)) val pc: Vec[PrunedAddr] = Vec(IBufferEnqueueWidth, PrunedAddr(VAddrBits))
val prevIBufEnqPtr: IBufPtr = new IBufPtr val prevIBufEnqPtr: IBufPtr = new IBufPtr

View File

@ -47,94 +47,57 @@ class IBufBankPtr(implicit p: Parameters) extends CircularQueuePtr[IBufBankPtr](
p(XSCoreParamsKey).frontendParameters.ibufferParameters.NumReadBank p(XSCoreParamsKey).frontendParameters.ibufferParameters.NumReadBank
) {} ) {}
// FIXME: this is to encode cross-page & ill into 3 bits, remove this when we support ill & hwe in ExceptionType
object IBufferExceptionType extends EnumUInt(8) {
def None: UInt = "b000".U(width.W)
def NonCrossPF: UInt = "b001".U(width.W)
def NonCrossGPF: UInt = "b010".U(width.W)
def NonCrossAF: UInt = "b011".U(width.W)
def RvcII: UInt = "b100".U(width.W) // illegal instruction
def CrossPF: UInt = "b101".U(width.W)
def CrossGPF: UInt = "b110".U(width.W)
def CrossAF: UInt = "b111".U(width.W)
def cvtFromFetchExcpAndCrossPageAndRVCII(fetchExcp: ExceptionType, crossPage: Bool, rvcIll: Bool): UInt =
MuxCase(
0.U,
Seq(
crossPage -> Cat(1.U(1.W), fetchExcp.value),
fetchExcp.hasException -> fetchExcp.value,
rvcIll -> this.RvcII
)
)
def isRVCII(uint: UInt): Bool = {
this.checkInputWidth(uint)
uint(2) && uint(1, 0) === 0.U
}
def isCrossPage(uint: UInt): Bool = {
this.checkInputWidth(uint)
uint(2) && uint(1, 0) =/= 0.U
}
def isPF(uint: UInt): Bool = uint(1, 0) === this.NonCrossPF(1, 0)
def isGPF(uint: UInt): Bool = uint(1, 0) === this.NonCrossGPF(1, 0)
def isAF(uint: UInt): Bool = uint(1, 0) === this.NonCrossAF(1, 0)
}
class IBufEntry(implicit p: Parameters) extends IBufferBundle { class IBufEntry(implicit p: Parameters) extends IBufferBundle {
val inst: UInt = UInt(32.W) val inst: UInt = UInt(32.W)
val pc: PrunedAddr = PrunedAddr(VAddrBits) val pc: PrunedAddr = PrunedAddr(VAddrBits)
val foldpc: UInt = UInt(MemPredPCWidth.W) val foldpc: UInt = UInt(MemPredPCWidth.W)
val pd: PreDecodeInfo = new PreDecodeInfo val pd: PreDecodeInfo = new PreDecodeInfo
val predTaken: Bool = Bool() val predTaken: Bool = Bool()
val identifiedCfi: Bool = Bool() val identifiedCfi: Bool = Bool()
val ftqPtr: FtqPtr = new FtqPtr val ftqPtr: FtqPtr = new FtqPtr
val instrEndOffset: UInt = UInt(FetchBlockInstOffsetWidth.W) val instrEndOffset: UInt = UInt(FetchBlockInstOffsetWidth.W)
val exceptionType: UInt = IBufferExceptionType() val exceptionType: ExceptionType = new ExceptionType
val backendException: Bool = Bool() val exceptionCrossPage: Bool = Bool()
val triggered: UInt = TriggerAction() val isBackendException: Bool = Bool()
val isLastInFtqEntry: Bool = Bool() val triggered: UInt = TriggerAction()
val isLastInFtqEntry: Bool = Bool()
val debug_seqNum: UInt = InstSeqNum() val debug_seqNum: UInt = InstSeqNum()
def fromFetch(fetch: FetchToIBuffer, i: Int): IBufEntry = { def fromFetch(fetch: FetchToIBuffer, i: Int): IBufEntry = {
inst := fetch.instrs(i) inst := fetch.instrs(i)
pc := fetch.pc(i) pc := fetch.pc(i)
foldpc := fetch.foldpc(i) foldpc := fetch.foldpc(i)
pd := fetch.pd(i) pd := fetch.pd(i)
predTaken := fetch.instrEndOffset(i).taken predTaken := fetch.instrEndOffset(i).taken
identifiedCfi := fetch.identifiedCfi(i) identifiedCfi := fetch.identifiedCfi(i)
ftqPtr := fetch.ftqPtr ftqPtr := fetch.ftqPtr
instrEndOffset := fetch.instrEndOffset(i).offset instrEndOffset := fetch.instrEndOffset(i).offset
exceptionType := IBufferExceptionType.cvtFromFetchExcpAndCrossPageAndRVCII( exceptionType := fetch.exceptionType(i)
fetch.exceptionType(i), exceptionCrossPage := fetch.exceptionCrossPage(i)
fetch.crossPageIPFFix(i), isBackendException := fetch.isBackendException(i)
fetch.illegalInstr(i) triggered := fetch.triggered(i)
) isLastInFtqEntry := fetch.isLastInFtqEntry(i)
backendException := fetch.backendException(i) debug_seqNum := fetch.debug_seqNum(i)
triggered := fetch.triggered(i)
isLastInFtqEntry := fetch.isLastInFtqEntry(i)
debug_seqNum := fetch.debug_seqNum(i)
this this
} }
def toIBufOutEntry: IBufOutEntry = { def toIBufOutEntry: IBufOutEntry = {
val result = Wire(new IBufOutEntry) val result = Wire(new IBufOutEntry)
result.inst := inst result.inst := inst
result.pc := pc result.pc := pc
result.foldpc := foldpc result.foldpc := foldpc
result.pd := pd result.pd := pd
result.predTaken := predTaken result.predTaken := predTaken
result.identifiedCfi := identifiedCfi result.identifiedCfi := identifiedCfi
result.ftqPtr := ftqPtr result.ftqPtr := ftqPtr
result.exceptionType := exceptionType result.exceptionType := exceptionType
result.backendException := backendException result.exceptionCrossPage := exceptionCrossPage
result.triggered := triggered result.isBackendException := isBackendException
result.isLastInFtqEntry := isLastInFtqEntry result.triggered := triggered
result.debug_seqNum := debug_seqNum result.isLastInFtqEntry := isLastInFtqEntry
result.instrEndOffset := instrEndOffset result.debug_seqNum := debug_seqNum
result.instrEndOffset := instrEndOffset
result result
} }
} }
@ -143,18 +106,19 @@ class IBufEntry(implicit p: Parameters) extends IBufferBundle {
// In the future, the backend will perform certain computations // In the future, the backend will perform certain computations
// in the IBuffer, which will be differentiated from IBufEntry. // in the IBuffer, which will be differentiated from IBufEntry.
class IBufOutEntry(implicit p: Parameters) extends IBufferBundle { class IBufOutEntry(implicit p: Parameters) extends IBufferBundle {
val inst: UInt = UInt(32.W) val inst: UInt = UInt(32.W)
val pc: PrunedAddr = PrunedAddr(VAddrBits) val pc: PrunedAddr = PrunedAddr(VAddrBits)
val foldpc: UInt = UInt(MemPredPCWidth.W) val foldpc: UInt = UInt(MemPredPCWidth.W)
val pd: PreDecodeInfo = new PreDecodeInfo val pd: PreDecodeInfo = new PreDecodeInfo
val predTaken: Bool = Bool() val predTaken: Bool = Bool()
val identifiedCfi: Bool = Bool() val identifiedCfi: Bool = Bool()
val ftqPtr: FtqPtr = new FtqPtr val ftqPtr: FtqPtr = new FtqPtr
val exceptionType: UInt = IBufferExceptionType() val exceptionType: ExceptionType = new ExceptionType
val backendException: Bool = Bool() val exceptionCrossPage: Bool = Bool()
val triggered: UInt = TriggerAction() val isBackendException: Bool = Bool()
val isLastInFtqEntry: Bool = Bool() val triggered: UInt = TriggerAction()
val instrEndOffset: UInt = UInt(FetchBlockInstOffsetWidth.W) val isLastInFtqEntry: Bool = Bool()
val instrEndOffset: UInt = UInt(FetchBlockInstOffsetWidth.W)
val debug_seqNum: UInt = InstSeqNum() val debug_seqNum: UInt = InstSeqNum()
@ -164,16 +128,17 @@ class IBufOutEntry(implicit p: Parameters) extends IBufferBundle {
cf.pc := pc.toUInt cf.pc := pc.toUInt
cf.foldpc := foldpc cf.foldpc := foldpc
cf.exceptionVec := 0.U.asTypeOf(ExceptionVec()) cf.exceptionVec := 0.U.asTypeOf(ExceptionVec())
cf.exceptionVec(ExceptionNO.instrPageFault) := IBufferExceptionType.isPF(exceptionType) cf.exceptionVec(ExceptionNO.instrPageFault) := exceptionType.isPf
cf.exceptionVec(ExceptionNO.instrGuestPageFault) := IBufferExceptionType.isGPF(exceptionType) cf.exceptionVec(ExceptionNO.instrGuestPageFault) := exceptionType.isGpf
cf.exceptionVec(ExceptionNO.instrAccessFault) := IBufferExceptionType.isAF(exceptionType) cf.exceptionVec(ExceptionNO.instrAccessFault) := exceptionType.isAf
cf.exceptionVec(ExceptionNO.illegalInstr) := IBufferExceptionType.isRVCII(exceptionType) cf.exceptionVec(ExceptionNO.illegalInstr) := exceptionType.isIll
cf.backendException := backendException cf.exceptionVec(ExceptionNO.hardwareError) := exceptionType.isHwe
cf.backendException := isBackendException
cf.trigger := triggered cf.trigger := triggered
cf.pd := pd cf.pd := pd
cf.pred_taken := predTaken cf.pred_taken := predTaken
cf.identifiedCfi := identifiedCfi cf.identifiedCfi := identifiedCfi
cf.crossPageIPFFix := IBufferExceptionType.isCrossPage(exceptionType) cf.crossPageIPFFix := exceptionCrossPage
cf.storeSetHit := DontCare cf.storeSetHit := DontCare
cf.waitForRobIdx := DontCare cf.waitForRobIdx := DontCare
cf.loadWaitBit := DontCare cf.loadWaitBit := DontCare

View File

@ -265,6 +265,7 @@ class MissRespBundle(implicit p: Parameters) extends ICacheBundle {
val waymask: UInt = UInt(nWays.W) val waymask: UInt = UInt(nWays.W)
val data: UInt = UInt(blockBits.W) val data: UInt = UInt(blockBits.W)
val corrupt: Bool = Bool() val corrupt: Bool = Bool()
val denied: Bool = Bool()
} }
/* ***** Mshr ***** */ /* ***** Mshr ***** */

View File

@ -232,13 +232,20 @@ class ICacheMainPipe(implicit p: Parameters) extends ICacheModule
) )
}) })
private val s1_l2Corrupt = VecInit((0 until PortNumber).map { i => private val s1_tlCorrupt = VecInit((0 until PortNumber).map { i =>
DataHoldBypass( DataHoldBypass(
s1_mshrHits(i) && fromMiss.bits.corrupt, s1_mshrHits(i) && fromMiss.bits.corrupt,
s1_mshrHits(i) || RegNext(s0_fire) s1_mshrHits(i) || RegNext(s0_fire)
) )
}) })
private val s1_tlDenied = VecInit((0 until PortNumber).map { i =>
DataHoldBypass(
s1_mshrHits(i) && fromMiss.bits.denied,
s1_mshrHits(i) || RegNext(s0_fire)
)
})
/* *** Update replacer *** */ /* *** Update replacer *** */
(0 until PortNumber).foreach { i => (0 until PortNumber).foreach { i =>
io.replacerTouch.req(i).bits.vSetIdx := s1_vSetIdx(i) io.replacerTouch.req(i).bits.vSetIdx := s1_vSetIdx(i)
@ -358,11 +365,13 @@ class ICacheMainPipe(implicit p: Parameters) extends ICacheModule
private val s1_fetchFinish = !s1_shouldFetch.reduce(_ || _) private val s1_fetchFinish = !s1_shouldFetch.reduce(_ || _)
// also raise af if l2 corrupt is detected // also raise af if l2 corrupt is detected
private val s1_l2Exception = ExceptionType(hasAf = s1_l2Corrupt.reduce(_ || _)) private val s1_tlException = (s1_tlCorrupt zip s1_tlDenied).map { case (corrupt, denied) =>
ExceptionType.fromTileLink(corrupt, denied, s1_valid) // s1_valid used only for assertion
}.reduce(_ || _)
// NOTE: do NOT raise af if meta/data corrupt is detected, they are automatically recovered by re-fetching from L2 // NOTE: do NOT raise af if meta/data corrupt is detected, they are automatically recovered by re-fetching from L2
// merge s2 exceptions, itlb/pmp has the highest priority, then l2 // merge s2 exceptions, itlb/pmp has the highest priority, then l2
private val s1_exceptionOut = s1_exception || s1_l2Exception private val s1_exceptionOut = s1_exception || s1_tlException
/* *** send response to IFU *** */ /* *** send response to IFU *** */
toIfu.valid := s1_fire toIfu.valid := s1_fire

View File

@ -182,19 +182,20 @@ class ICacheMissUnit(edge: TLEdgeOut)(implicit p: Parameters) extends ICacheModu
private val lastFireNext = RegNext(lastFire) private val lastFireNext = RegNext(lastFire)
private val idNext = RegNext(io.memGrant.bits.source) private val idNext = RegNext(io.memGrant.bits.source)
// if any beat is corrupt, the whole response (to mainPipe/metaArray/dataArray) is corrupt
private val corruptReg = RegInit(false.B) private val corruptReg = RegInit(false.B)
when(io.memGrant.fire && edge.hasData(io.memGrant.bits) && io.memGrant.bits.corrupt) { private val deniedReg = RegInit(false.B)
// Set corruptReg when any beat is corrupt when(io.memGrant.fire && edge.hasData(io.memGrant.bits)) {
// This is actually when(xxx.fire && xxx.hasData) { corruptReg := corruptReg || io.mem_grant.bits.corrupt } // Set corruptReg / deniedReg when any beat is corrupt / denied
corruptReg := true.B corruptReg := corruptReg || io.memGrant.bits.corrupt
deniedReg := deniedReg || io.memGrant.bits.denied
}.elsewhen(lastFireNext) { }.elsewhen(lastFireNext) {
// Clear corruptReg when response it sent to mainPipe // Clear corruptReg / deniedReg when response it sent to mainPipe
// This used to be io.resp.valid (lastFireNext && mshrValid) but when mshr is flushed by io.flush/fencei, // This used to be io.resp.valid (lastFireNext && mshrValid) but when mshr is flushed by io.flush/fencei,
// mshrValid is false.B and corruptReg will never be cleared, that's not correct // mshrValid is false.B and corruptReg will never be cleared, that's not correct
// so we remove mshrValid here, and the condition leftover is lastFireNext // so we remove mshrValid here, and the condition leftover is lastFireNext
// or, actually, io.resp.valid || (lastFireNext && !mshrValid) // or, actually, io.resp.valid || (lastFireNext && !mshrValid)
corruptReg := false.B corruptReg := false.B
deniedReg := false.B
} }
/** /**
@ -225,6 +226,7 @@ class ICacheMissUnit(edge: TLEdgeOut)(implicit p: Parameters) extends ICacheModu
// unnecessary response will be dropped by mainPipe/prefetchPipe/wayLookup since their sx_valid is set to false // unnecessary response will be dropped by mainPipe/prefetchPipe/wayLookup since their sx_valid is set to false
private val respValid = mshrValid && lastFireNext private val respValid = mshrValid && lastFireNext
// NOTE: but we should not write meta/dataArray when flush/fencei // NOTE: but we should not write meta/dataArray when flush/fencei
// NOTE: tilelink spec asks corrupt to be set when denied is set, so we don't need to check deniedReg here
private val writeSramValid = respValid && !corruptReg && !io.flush && !io.fencei private val writeSramValid = respValid && !corruptReg && !io.flush && !io.fencei
// write SRAM // write SRAM
@ -253,6 +255,7 @@ class ICacheMissUnit(edge: TLEdgeOut)(implicit p: Parameters) extends ICacheModu
io.resp.bits.waymask := waymask io.resp.bits.waymask := waymask
io.resp.bits.data := respDataReg.asUInt io.resp.bits.data := respDataReg.asUInt
io.resp.bits.corrupt := corruptReg io.resp.bits.corrupt := corruptReg
io.resp.bits.denied := deniedReg
// we are safe to enter wfi if all entries have no pending response from L2 // we are safe to enter wfi if all entries have no pending response from L2
io.wfi.wfiSafe := allMshr.map(_.io.wfi.wfiSafe).reduce(_ && _) io.wfi.wfiSafe := allMshr.map(_.io.wfi.wfiSafe).reduce(_ && _)

View File

@ -509,7 +509,9 @@ class Ifu(implicit p: Parameters) extends IfuModule
private val s4_alignExpdInstr = VecInit(rvcExpanders.map { expander: RvcExpander => private val s4_alignExpdInstr = VecInit(rvcExpanders.map { expander: RvcExpander =>
Mux(expander.io.ill, expander.io.in, expander.io.out.bits) Mux(expander.io.ill, expander.io.in, expander.io.out.bits)
}) })
private val s4_alignIll = VecInit(rvcExpanders.map(_.io.ill)) private val s4_alignRvcException = VecInit(rvcExpanders.map { expander: RvcExpander =>
ExceptionType.fromRvcExpander(expander.io.ill)
})
private val s4_alignPdWire = RegEnable(s3_alignPd, s3_fire) private val s4_alignPdWire = RegEnable(s3_alignPd, s3_fire)
private val s4_alignPds = WireInit(s4_alignPdWire) private val s4_alignPds = WireInit(s4_alignPdWire)
@ -678,19 +680,23 @@ class Ifu(implicit p: Parameters) extends IfuModule
// mark the exception only on first instruction // mark the exception only on first instruction
// TODO: store only the first exception in IBuffer, instead of store in every entry // TODO: store only the first exception in IBuffer, instead of store in every entry
io.toIBuffer.bits.exceptionType := (0 until IBufferEnqueueWidth).map { io.toIBuffer.bits.exceptionType := (0 until IBufferEnqueueWidth).map {
i => Mux(i.U === s4_prevIBufEnqPtr.value(1, 0), s4_icacheMeta(0).exception, ExceptionType.None) i =>
Mux(
i.U === s4_prevIBufEnqPtr.value(1, 0),
s4_icacheMeta(0).exception,
ExceptionType.None
) || s4_alignRvcException(i)
} }
// backendException only needs to be set for the first instruction. // backendException only needs to be set for the first instruction.
// Other instructions in the same block may have pf or af set, // Other instructions in the same block may have pf or af set,
// which is a side effect of the first instruction and actually not necessary. // which is a side effect of the first instruction and actually not necessary.
io.toIBuffer.bits.backendException := (0 until IBufferEnqueueWidth).map { io.toIBuffer.bits.isBackendException := (0 until IBufferEnqueueWidth).map {
i => i.U === s4_prevIBufEnqPtr.value(1, 0) && s4_icacheMeta(0).isBackendException i => i.U === s4_prevIBufEnqPtr.value(1, 0) && s4_icacheMeta(0).isBackendException
} }
// if we have last half RV-I instruction, and has exception, we need to tell backend to caculate the correct pc // if we have last half RV-I instruction, and has exception, we need to tell backend to caculate the correct pc
io.toIBuffer.bits.crossPageIPFFix := (0 until IBufferEnqueueWidth).map { io.toIBuffer.bits.exceptionCrossPage := (0 until IBufferEnqueueWidth).map {
i => i.U === s4_prevIBufEnqPtr.value(1, 0) && s4_icacheMeta(0).exception.hasException && s4_prevLastIsHalfRvi i => i.U === s4_prevIBufEnqPtr.value(1, 0) && s4_icacheMeta(0).exception.hasException && s4_prevLastIsHalfRvi
} }
io.toIBuffer.bits.illegalInstr := s4_alignIll
io.toIBuffer.bits.triggered := s4_alignTriggered io.toIBuffer.bits.triggered := s4_alignTriggered
io.toIBuffer.bits.identifiedCfi := s4_alignCompactInfo.identifiedCfi io.toIBuffer.bits.identifiedCfi := s4_alignCompactInfo.identifiedCfi
@ -742,6 +748,9 @@ class Ifu(implicit p: Parameters) extends IfuModule
when(s4_reqIsUncache) { when(s4_reqIsUncache) {
val inst = s4_uncacheData val inst = s4_uncacheData
val (brType, isCall, isRet) = getBrInfo(inst) val (brType, isCall, isRet) = getBrInfo(inst)
val uncacheRvcException = ExceptionType.fromRvcExpander(uncacheRvcExpander.io.ill)
io.toIBuffer.bits.instrs(s4_shiftNum) := Mux( io.toIBuffer.bits.instrs(s4_shiftNum) := Mux(
uncacheRvcExpander.io.ill, uncacheRvcExpander.io.ill,
uncacheRvcExpander.io.in, uncacheRvcExpander.io.in,
@ -756,11 +765,11 @@ class Ifu(implicit p: Parameters) extends IfuModule
io.toIBuffer.bits.pd(s4_shiftNum).isRet := isRet io.toIBuffer.bits.pd(s4_shiftNum).isRet := isRet
io.toIBuffer.bits.instrEndOffset(s4_shiftNum).offset := Mux(prevUncacheCrossPage || uncacheIsRvc, 0.U, 1.U) io.toIBuffer.bits.instrEndOffset(s4_shiftNum).offset := Mux(prevUncacheCrossPage || uncacheIsRvc, 0.U, 1.U)
io.toIBuffer.bits.exceptionType(s4_shiftNum) := uncacheException io.toIBuffer.bits.exceptionType(s4_shiftNum) := uncacheException || uncacheRvcException
// execption can happen in next page only when cross page. // execption can happen in next page only when cross page.
io.toIBuffer.bits.crossPageIPFFix(s4_shiftNum) := prevUncacheCrossPage && uncacheException.hasException io.toIBuffer.bits.exceptionCrossPage(s4_shiftNum) := prevUncacheCrossPage && uncacheException.hasException
io.toIBuffer.bits.illegalInstr(s4_shiftNum) := uncacheRvcExpander.io.ill
io.toIBuffer.bits.enqEnable := s4_alignBlockStartPos.asUInt io.toIBuffer.bits.enqEnable := s4_alignBlockStartPos.asUInt
uncacheFlushWb.bits.isRVC := uncacheIsRvc uncacheFlushWb.bits.isRVC := uncacheIsRvc
uncacheFlushWb.bits.attribute := BranchAttribute(brType, Cat(isCall, isRet)) uncacheFlushWb.bits.attribute := BranchAttribute(brType, Cat(isCall, isRet))

View File

@ -114,7 +114,7 @@ class IfuUncacheUnit(implicit p: Parameters) extends IfuModule with IfuHelper {
is(UncacheFsmState.WaitResp) { is(UncacheFsmState.WaitResp) {
when(fromUncache.fire) { when(fromUncache.fire) {
val exception = ExceptionType(hasAf = fromUncache.bits.corrupt) val exception = ExceptionType.fromTileLink(fromUncache.bits.corrupt, fromUncache.bits.denied)
val crossPage = fromUncache.bits.incomplete val crossPage = fromUncache.bits.incomplete
uncacheState := UncacheFsmState.Idle uncacheState := UncacheFsmState.Idle
uncacheException := exception uncacheException := exception

View File

@ -26,7 +26,9 @@ class InstrUncacheReq(implicit p: Parameters) extends InstrUncacheBundle {
} }
class InstrUncacheResp(implicit p: Parameters) extends InstrUncacheBundle { class InstrUncacheResp(implicit p: Parameters) extends InstrUncacheBundle {
val data: UInt = UInt(32.W) // TODO: add a const for InstrLen, maybe in XSParameters, and use it all over the repo // TODO: add a const for InstrLen, maybe in XSParameters, and use it all over the repo, remove magic number 32
val corrupt: Bool = Bool() val data: UInt = UInt(32.W)
val corrupt: Bool = Bool()
val denied: Bool = Bool()
val incomplete: Bool = Bool() // whether this.data is incomplete (e.g. crossing a page boundary) val incomplete: Bool = Bool() // whether this.data is incomplete (e.g. crossing a page boundary)
} }

View File

@ -94,11 +94,13 @@ class InstrUncacheEntry(edge: TLEdgeOut)(implicit p: Parameters) extends InstrUn
private val respDataReg = RegInit(VecInit.fill(2)(0.U(16.W))) // FIXME: 2 * rvc, how to avoid magic number? private val respDataReg = RegInit(VecInit.fill(2)(0.U(16.W))) // FIXME: 2 * rvc, how to avoid magic number?
private val respCorruptReg = RegInit(false.B) private val respCorruptReg = RegInit(false.B)
private val respDeniedReg = RegInit(false.B)
// send response to InstrUncache // send response to InstrUncache
io.resp.valid := state === State.SendResp && !needFlush io.resp.valid := state === State.SendResp && !needFlush
io.resp.bits.data := respDataReg.asUInt io.resp.bits.data := respDataReg.asUInt
io.resp.bits.corrupt := respCorruptReg io.resp.bits.corrupt := respCorruptReg
io.resp.bits.denied := respDeniedReg
io.resp.bits.incomplete := crossPageBoundary io.resp.bits.incomplete := crossPageBoundary
// state transfer // state transfer
@ -149,6 +151,7 @@ class InstrUncacheEntry(edge: TLEdgeOut)(implicit p: Parameters) extends InstrUn
respDataReg(1) := shiftedBusData(31, 16) respDataReg(1) := shiftedBusData(31, 16)
} }
respCorruptReg := io.mmioGrant.bits.corrupt respCorruptReg := io.mmioGrant.bits.corrupt
respDeniedReg := io.mmioGrant.bits.denied
} }
} }