1
0

Refactor package hierarchy and remove legacy bus protocol implementations (#845)

* Refactors package hierarchy.

Additionally:
  - Removes legacy ground tests and configs
  - Removes legacy bus protocol implementations
  - Removes NTiles
  - Adds devices package
  - Adds more functions to util package
This commit is contained in:
Henry Cook
2017-07-07 10:48:16 -07:00
committed by GitHub
parent c28c23150d
commit 4c595d175c
238 changed files with 1347 additions and 10978 deletions

View File

@ -0,0 +1,35 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
import freechips.rocketchip.util.GenericParameterizedBundle
abstract class AHBBundleBase(params: AHBBundleParameters) extends GenericParameterizedBundle(params)
// Signal directions are from the master's point-of-view
class AHBBundle(params: AHBBundleParameters) extends AHBBundleBase(params)
{
// Flow control signals from the master
val hmastlock = Bool(OUTPUT)
val htrans = UInt(OUTPUT, width = params.transBits)
val hsel = Bool(OUTPUT)
val hready = Bool(OUTPUT) // on a master, drive this from readyout
// Payload signals
val hwrite = Bool(OUTPUT)
val haddr = UInt(OUTPUT, width = params.addrBits)
val hsize = UInt(OUTPUT, width = params.sizeBits)
val hburst = UInt(OUTPUT, width = params.burstBits)
val hprot = UInt(OUTPUT, width = params.protBits)
val hwdata = UInt(OUTPUT, width = params.dataBits)
val hreadyout = Bool(INPUT)
val hresp = Bool(INPUT)
val hrdata = UInt(INPUT, width = params.dataBits)
}
object AHBBundle
{
def apply(params: AHBBundleParameters) = new AHBBundle(params)
}

View File

@ -0,0 +1,47 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
import chisel3.internal.sourceinfo.SourceInfo
import freechips.rocketchip.diplomacy._
object AHBImp extends NodeImp[AHBMasterPortParameters, AHBSlavePortParameters, AHBEdgeParameters, AHBEdgeParameters, AHBBundle]
{
def edgeO(pd: AHBMasterPortParameters, pu: AHBSlavePortParameters): AHBEdgeParameters = AHBEdgeParameters(pd, pu)
def edgeI(pd: AHBMasterPortParameters, pu: AHBSlavePortParameters): AHBEdgeParameters = AHBEdgeParameters(pd, pu)
def bundleO(eo: AHBEdgeParameters): AHBBundle = AHBBundle(eo.bundle)
def bundleI(ei: AHBEdgeParameters): AHBBundle = AHBBundle(ei.bundle)
def colour = "#00ccff" // bluish
override def labelI(ei: AHBEdgeParameters) = (ei.slave.beatBytes * 8).toString
override def labelO(eo: AHBEdgeParameters) = (eo.slave.beatBytes * 8).toString
override def mixO(pd: AHBMasterPortParameters, node: OutwardNode[AHBMasterPortParameters, AHBSlavePortParameters, AHBBundle]): AHBMasterPortParameters =
pd.copy(masters = pd.masters.map { c => c.copy (nodePath = node +: c.nodePath) })
override def mixI(pu: AHBSlavePortParameters, node: InwardNode[AHBMasterPortParameters, AHBSlavePortParameters, AHBBundle]): AHBSlavePortParameters =
pu.copy(slaves = pu.slaves.map { m => m.copy (nodePath = node +: m.nodePath) })
}
// Nodes implemented inside modules
case class AHBIdentityNode() extends IdentityNode(AHBImp)
case class AHBMasterNode(portParams: Seq[AHBMasterPortParameters]) extends SourceNode(AHBImp)(portParams)
case class AHBSlaveNode(portParams: Seq[AHBSlavePortParameters]) extends SinkNode(AHBImp)(portParams)
case class AHBNexusNode(
masterFn: Seq[AHBMasterPortParameters] => AHBMasterPortParameters,
slaveFn: Seq[AHBSlavePortParameters] => AHBSlavePortParameters,
numMasterPorts: Range.Inclusive = 1 to 999,
numSlavePorts: Range.Inclusive = 1 to 999)
extends NexusNode(AHBImp)(masterFn, slaveFn, numMasterPorts, numSlavePorts)
// Nodes passed from an inner module
case class AHBOutputNode() extends OutputNode(AHBImp)
case class AHBInputNode() extends InputNode(AHBImp)
// Nodes used for external ports
case class AHBBlindOutputNode(portParams: Seq[AHBSlavePortParameters]) extends BlindOutputNode(AHBImp)(portParams)
case class AHBBlindInputNode(portParams: Seq[AHBMasterPortParameters]) extends BlindInputNode(AHBImp)(portParams)
case class AHBInternalOutputNode(portParams: Seq[AHBSlavePortParameters]) extends InternalOutputNode(AHBImp)(portParams)
case class AHBInternalInputNode(portParams: Seq[AHBMasterPortParameters]) extends InternalInputNode(AHBImp)(portParams)

View File

@ -0,0 +1,96 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
import freechips.rocketchip.diplomacy._
import scala.math.max
case class AHBSlaveParameters(
address: Seq[AddressSet],
resources: Seq[Resource] = Nil,
regionType: RegionType.T = RegionType.GET_EFFECTS,
executable: Boolean = false, // processor can execute from this memory
nodePath: Seq[BaseNode] = Seq(),
supportsWrite: TransferSizes = TransferSizes.none,
supportsRead: TransferSizes = TransferSizes.none)
{
address.foreach { a => require (a.finite) }
address.combinations(2).foreach { case Seq(x,y) => require (!x.overlaps(y)) }
val name = nodePath.lastOption.map(_.lazyModule.name).getOrElse("disconnected")
val maxTransfer = max(supportsWrite.max, supportsRead.max)
val maxAddress = address.map(_.max).max
val minAlignment = address.map(_.alignment).min
// The device had better not support a transfer larger than it's alignment
require (minAlignment >= maxTransfer)
}
case class AHBSlavePortParameters(
slaves: Seq[AHBSlaveParameters],
beatBytes: Int)
{
require (!slaves.isEmpty)
require (isPow2(beatBytes))
val maxTransfer = slaves.map(_.maxTransfer).max
val maxAddress = slaves.map(_.maxAddress).max
// Check the link is not pointlessly wide
require (maxTransfer >= beatBytes)
// Check that the link can be implemented in AHB
require (maxTransfer <= beatBytes * AHBParameters.maxTransfer)
// Require disjoint ranges for addresses
slaves.combinations(2).foreach { case Seq(x,y) =>
x.address.foreach { a => y.address.foreach { b =>
require (!a.overlaps(b))
} }
}
}
case class AHBMasterParameters(
name: String,
nodePath: Seq[BaseNode] = Seq())
case class AHBMasterPortParameters(
masters: Seq[AHBMasterParameters])
case class AHBBundleParameters(
addrBits: Int,
dataBits: Int)
{
require (dataBits >= 8)
require (addrBits >= 1)
require (isPow2(dataBits))
// Bring the globals into scope
val transBits = AHBParameters.transBits
val burstBits = AHBParameters.burstBits
val protBits = AHBParameters.protBits
val sizeBits = AHBParameters.sizeBits
def union(x: AHBBundleParameters) =
AHBBundleParameters(
max(addrBits, x.addrBits),
max(dataBits, x.dataBits))
}
object AHBBundleParameters
{
val emptyBundleParams = AHBBundleParameters(addrBits = 1, dataBits = 8)
def union(x: Seq[AHBBundleParameters]) = x.foldLeft(emptyBundleParams)((x,y) => x.union(y))
def apply(master: AHBMasterPortParameters, slave: AHBSlavePortParameters) =
new AHBBundleParameters(
addrBits = log2Up(slave.maxAddress+1),
dataBits = slave.beatBytes * 8)
}
case class AHBEdgeParameters(
master: AHBMasterPortParameters,
slave: AHBSlavePortParameters)
{
val bundle = AHBBundleParameters(master, slave)
}

View File

@ -0,0 +1,39 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
object AHBParameters
{
// These are all fixed by the AHB standard:
val transBits = 2
val burstBits = 3
val protBits = 4
val sizeBits = 3 // 8*2^s
def TRANS_IDLE = UInt(0, width = transBits) // No transfer requested, not in a burst
def TRANS_BUSY = UInt(1, width = transBits) // No transfer requested, in a burst
def TRANS_NONSEQ = UInt(2, width = transBits) // First (potentially only) request in a burst
def TRANS_SEQ = UInt(3, width = transBits) // Following requests in a burst
def BURST_SINGLE = UInt(0, width = burstBits) // Single access (no burst)
def BURST_INCR = UInt(1, width = burstBits) // Incrementing burst of arbitrary length, not crossing 1KB
def BURST_WRAP4 = UInt(2, width = burstBits) // 4-beat wrapping burst
def BURST_INCR4 = UInt(3, width = burstBits) // 4-beat incrementing burst
def BURST_WRAP8 = UInt(4, width = burstBits) // 8-beat wrapping burst
def BURST_INCR8 = UInt(5, width = burstBits) // 8-beat incrementing burst
def BURST_WRAP16 = UInt(6, width = burstBits) // 16-beat wrapping burst
def BURST_INCR16 = UInt(7, width = burstBits) // 16-beat incrementing burst
val maxTransfer = 16
def RESP_OKAY = Bool(false)
def RESP_ERROR = Bool(true)
def PROT_DATA = UInt(1, width = protBits)
def PROT_PRIVILEDGED = UInt(2, width = protBits)
def PROT_BUFFERABLE = UInt(4, width = protBits)
def PROT_CACHEABLE = UInt(8, width = protBits)
def PROT_DEFAULT = PROT_DATA | PROT_PRIVILEDGED
}

View File

@ -0,0 +1,114 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
import freechips.rocketchip.config.Parameters
import freechips.rocketchip.diplomacy._
import freechips.rocketchip.regmapper._
import freechips.rocketchip.tilelink.{IntSourceNode, IntSourcePortSimple}
import freechips.rocketchip.util.{HeterogeneousBag, MaskGen}
import scala.math.{min,max}
class AHBRegisterNode(address: AddressSet, concurrency: Int = 0, beatBytes: Int = 4, undefZero: Boolean = true, executable: Boolean = false)
extends AHBSlaveNode(Seq(AHBSlavePortParameters(
Seq(AHBSlaveParameters(
address = Seq(address),
executable = executable,
supportsWrite = TransferSizes(1, min(address.alignment.toInt, beatBytes * AHBParameters.maxTransfer)),
supportsRead = TransferSizes(1, min(address.alignment.toInt, beatBytes * AHBParameters.maxTransfer)))),
beatBytes = beatBytes)))
{
require (address.contiguous)
// Calling this method causes the matching AHB bundle to be
// configured to route all requests to the listed RegFields.
def regmap(mapping: RegField.Map*) = {
val ahb = bundleIn(0)
val indexBits = log2Up((address.mask+1)/beatBytes)
val params = RegMapperParams(indexBits, beatBytes, 1)
val in = Wire(Decoupled(new RegMapperInput(params)))
val out = RegMapper(beatBytes, concurrency, undefZero, in, mapping:_*)
val d_phase = RegInit(Bool(false))
val d_taken = Reg(Bool())
val d_read = Reg(Bool())
val d_index = Reg(UInt(width = indexBits))
val d_mask = Reg(UInt(width = beatBytes))
// Only send the request to the RR once
d_taken := d_phase && in.ready
in.valid := d_phase && !d_taken
in.bits.read := d_read
in.bits.index := d_index
in.bits.data := ahb.hwdata
in.bits.mask := d_mask
in.bits.extra := UInt(0)
when (ahb.hready) { d_phase := Bool(false) }
ahb.hreadyout := !d_phase || out.valid
ahb.hresp := AHBParameters.RESP_OKAY
ahb.hrdata := out.bits.data
val request = ahb.htrans === AHBParameters.TRANS_NONSEQ || ahb.htrans === AHBParameters.TRANS_SEQ
when (ahb.hready && ahb.hsel && request) {
assert (!in.valid || in.ready)
d_phase := Bool(true)
d_taken := Bool(false)
d_read := !ahb.hwrite
d_index := ahb.haddr >> log2Ceil(beatBytes)
d_mask := MaskGen(ahb.haddr, ahb.hsize, beatBytes)
}
out.ready := Bool(true)
assert (d_phase || !out.valid)
}
}
object AHBRegisterNode
{
def apply(address: AddressSet, concurrency: Int = 0, beatBytes: Int = 4, undefZero: Boolean = true, executable: Boolean = false) =
new AHBRegisterNode(address, concurrency, beatBytes, undefZero, executable)
}
// These convenience methods below combine to make it possible to create a AHB
// register mapped device from a totally abstract register mapped device.
abstract class AHBRegisterRouterBase(address: AddressSet, interrupts: Int, concurrency: Int, beatBytes: Int, undefZero: Boolean, executable: Boolean)(implicit p: Parameters) extends LazyModule
{
val node = AHBRegisterNode(address, concurrency, beatBytes, undefZero, executable)
val intnode = IntSourceNode(IntSourcePortSimple(num = interrupts))
}
case class AHBRegBundleArg(interrupts: HeterogeneousBag[Vec[Bool]], in: HeterogeneousBag[AHBBundle])(implicit val p: Parameters)
class AHBRegBundleBase(arg: AHBRegBundleArg) extends Bundle
{
implicit val p = arg.p
val interrupts = arg.interrupts
val in = arg.in
}
class AHBRegBundle[P](val params: P, arg: AHBRegBundleArg) extends AHBRegBundleBase(arg)
class AHBRegModule[P, B <: AHBRegBundleBase](val params: P, bundleBuilder: => B, router: AHBRegisterRouterBase)
extends LazyModuleImp(router) with HasRegMap
{
val io = bundleBuilder
val interrupts = if (io.interrupts.isEmpty) Vec(0, Bool()) else io.interrupts(0)
def regmap(mapping: RegField.Map*) = router.node.regmap(mapping:_*)
}
class AHBRegisterRouter[B <: AHBRegBundleBase, M <: LazyModuleImp]
(val base: BigInt, val interrupts: Int = 0, val size: BigInt = 4096, val concurrency: Int = 0, val beatBytes: Int = 4, undefZero: Boolean = true, executable: Boolean = false)
(bundleBuilder: AHBRegBundleArg => B)
(moduleBuilder: (=> B, AHBRegisterRouterBase) => M)(implicit p: Parameters)
extends AHBRegisterRouterBase(AddressSet(base, size-1), interrupts, concurrency, beatBytes, undefZero, executable)
{
require (isPow2(size))
// require (size >= 4096) ... not absolutely required, but highly recommended
lazy val module = moduleBuilder(bundleBuilder(AHBRegBundleArg(intnode.bundleOut, node.bundleIn)), this)
}

View File

@ -0,0 +1,102 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
import freechips.rocketchip.config.Parameters
import freechips.rocketchip.diplomacy._
import freechips.rocketchip.util._
class AHBRAM(address: AddressSet, executable: Boolean = true, beatBytes: Int = 4, fuzzHreadyout: Boolean = false)(implicit p: Parameters) extends LazyModule
{
val node = AHBSlaveNode(Seq(AHBSlavePortParameters(
Seq(AHBSlaveParameters(
address = List(address),
regionType = RegionType.UNCACHED,
executable = executable,
supportsRead = TransferSizes(1, beatBytes * AHBParameters.maxTransfer),
supportsWrite = TransferSizes(1, beatBytes * AHBParameters.maxTransfer))),
beatBytes = beatBytes)))
// We require the address range to include an entire beat (for the write mask)
require ((address.mask & (beatBytes-1)) == beatBytes-1)
lazy val module = new LazyModuleImp(this) {
val io = new Bundle {
val in = node.bundleIn
}
def bigBits(x: BigInt, tail: List[Boolean] = List.empty[Boolean]): List[Boolean] =
if (x == 0) tail.reverse else bigBits(x >> 1, ((x & 1) == 1) :: tail)
val mask = bigBits(address.mask >> log2Ceil(beatBytes))
val in = io.in(0)
// The mask and address during the address phase
val a_access = in.htrans === AHBParameters.TRANS_NONSEQ || in.htrans === AHBParameters.TRANS_SEQ
val a_request = in.hready && in.hsel && a_access
val a_mask = MaskGen(in.haddr, in.hsize, beatBytes)
val a_address = Cat((mask zip (in.haddr >> log2Ceil(beatBytes)).toBools).filter(_._1).map(_._2).reverse)
val a_write = in.hwrite
// The data phase signals
val d_wdata = Vec.tabulate(beatBytes) { i => in.hwdata(8*(i+1)-1, 8*i) }
// AHB writes must occur during the data phase; this poses a structural
// hazard with reads which must occur during the address phase. To solve
// this problem, we delay the writes until there is a free cycle.
//
// The idea is to record the address information from address phase and
// then as soon as possible flush the pending write. This cannot be done
// on a cycle when there is an address phase read, but on any other cycle
// the write will execute. In the case of reads following a write, the
// result must bypass data from the pending write into the read if they
// happen to have matching address.
// Pending write?
val p_valid = RegInit(Bool(false))
val p_address = Reg(a_address)
val p_mask = Reg(a_mask)
val p_latch_d = Reg(Bool())
val p_wdata = d_wdata holdUnless p_latch_d
// Use single-ported memory with byte-write enable
val mem = SeqMem(1 << mask.filter(b=>b).size, Vec(beatBytes, Bits(width = 8)))
// Decide if the SRAM port is used for reading or (potentially) writing
val read = a_request && !a_write
// In case we choose to stall, we need to hold the read data
val d_rdata = mem.readAndHold(a_address, read)
// Whenever the port is not needed for reading, execute pending writes
when (!read && p_valid) {
p_valid := Bool(false)
mem.write(p_address, p_wdata, p_mask.toBools)
}
// Record the request for later?
p_latch_d := a_request && a_write
when (a_request && a_write) {
p_valid := Bool(true)
p_address := a_address
p_mask := a_mask
}
// Does the read need to be muxed with the previous write?
val a_bypass = a_address === p_address && p_valid
val d_bypass = RegEnable(a_bypass, a_request)
// Mux in data from the pending write
val muxdata = Vec((p_mask.toBools zip (p_wdata zip d_rdata))
map { case (m, (p, r)) => Mux(d_bypass && m, p, r) })
// Don't fuzz hready when not in data phase
val d_request = Reg(Bool(false))
when (in.hready) { d_request := Bool(false) }
when (a_request) { d_request := Bool(true) }
// Finally, the outputs
in.hreadyout := (if(fuzzHreadyout) { !d_request || LFSR16(Bool(true))(0) } else { Bool(true) })
in.hresp := AHBParameters.RESP_OKAY
in.hrdata := Mux(in.hreadyout, muxdata.asUInt, UInt(0))
}
}

View File

@ -0,0 +1,102 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
import freechips.rocketchip.config.Parameters
import freechips.rocketchip.devices.tilelink.TLTestRAM
import freechips.rocketchip.diplomacy._
import freechips.rocketchip.tilelink._
import freechips.rocketchip.unittest._
class RRTest0(address: BigInt)(implicit p: Parameters) extends AHBRegisterRouter(address, 0, 32, 0, 4)(
new AHBRegBundle((), _) with RRTest0Bundle)(
new AHBRegModule((), _, _) with RRTest0Module)
class RRTest1(address: BigInt)(implicit p: Parameters) extends AHBRegisterRouter(address, 0, 32, 1, 4, false)(
new AHBRegBundle((), _) with RRTest1Bundle)(
new AHBRegModule((), _, _) with RRTest1Module)
class AHBFuzzNative(aFlow: Boolean, txns: Int)(implicit p: Parameters) extends LazyModule
{
val fuzz = LazyModule(new TLFuzzer(txns))
val model = LazyModule(new TLRAMModel("AHBFuzzNative"))
var xbar = LazyModule(new AHBFanout)
val ram = LazyModule(new AHBRAM(AddressSet(0x0, 0xff)))
val gpio = LazyModule(new RRTest0(0x100))
model.node := fuzz.node
xbar.node := TLToAHB(aFlow)(TLDelayer(0.1)(model.node))
ram.node := xbar.node
gpio.node := xbar.node
lazy val module = new LazyModuleImp(this) with HasUnitTestIO {
io.finished := fuzz.module.io.finished
}
}
class AHBNativeTest(aFlow: Boolean, txns: Int = 5000, timeout: Int = 500000)(implicit p: Parameters) extends UnitTest(timeout) {
val dut = Module(LazyModule(new AHBFuzzNative(aFlow, txns)).module)
io.finished := dut.io.finished
}
class AHBFuzzMaster(aFlow: Boolean, txns: Int)(implicit p: Parameters) extends LazyModule
{
val node = AHBOutputNode()
val fuzz = LazyModule(new TLFuzzer(txns))
val model = LazyModule(new TLRAMModel("AHBFuzzMaster"))
model.node := fuzz.node
node :=
TLToAHB(aFlow)(
TLDelayer(0.2)(
TLBuffer(BufferParams.flow)(
TLDelayer(0.2)(
model.node))))
lazy val module = new LazyModuleImp(this) {
val io = new Bundle {
val out = node.bundleOut
val finished = Bool(OUTPUT)
}
io.finished := fuzz.module.io.finished
}
}
class AHBFuzzSlave()(implicit p: Parameters) extends LazyModule
{
val node = AHBInputNode()
val ram = LazyModule(new TLTestRAM(AddressSet(0x0, 0xfff)))
ram.node :=
TLFragmenter(4, 16)(
TLDelayer(0.2)(
TLBuffer(BufferParams.flow)(
TLDelayer(0.2)(
AHBToTL()(
node)))))
lazy val module = new LazyModuleImp(this) {
val io = new Bundle {
val in = node.bundleIn
}
}
}
class AHBFuzzBridge(aFlow: Boolean, txns: Int)(implicit p: Parameters) extends LazyModule
{
val master = LazyModule(new AHBFuzzMaster(aFlow, txns))
val slave = LazyModule(new AHBFuzzSlave)
slave.node := master.node
lazy val module = new LazyModuleImp(this) with HasUnitTestIO {
io.finished := master.module.io.finished
}
}
class AHBBridgeTest(aFlow: Boolean, txns: Int = 5000, timeout: Int = 500000)(implicit p: Parameters) extends UnitTest(timeout) {
val dut = Module(LazyModule(new AHBFuzzBridge(aFlow, txns)).module)
io.finished := dut.io.finished
}

View File

@ -0,0 +1,145 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
import chisel3.internal.sourceinfo.SourceInfo
import freechips.rocketchip.config.Parameters
import freechips.rocketchip.diplomacy._
import freechips.rocketchip.tilelink._
import freechips.rocketchip.util.MaskGen
case class AHBToTLNode() extends MixedAdapterNode(AHBImp, TLImp)(
dFn = { case AHBMasterPortParameters(masters) =>
TLClientPortParameters(clients = masters.map { m =>
TLClientParameters(name = m.name, nodePath = m.nodePath)
})
},
uFn = { mp => AHBSlavePortParameters(
slaves = mp.managers.map { m =>
def adjust(x: TransferSizes) = {
if (x.contains(mp.beatBytes)) {
TransferSizes(x.min, m.minAlignment.min(mp.beatBytes * AHBParameters.maxTransfer).toInt)
} else { // larger than beatBytes requires beatBytes if misaligned
x.intersect(TransferSizes(1, mp.beatBytes))
}
}
AHBSlaveParameters(
address = m.address,
resources = m.resources,
regionType = m.regionType,
executable = m.executable,
nodePath = m.nodePath,
supportsWrite = adjust(m.supportsPutFull),
supportsRead = adjust(m.supportsGet))},
beatBytes = mp.beatBytes)
})
class AHBToTL()(implicit p: Parameters) extends LazyModule
{
val node = AHBToTLNode()
lazy val module = new LazyModuleImp(this) {
val io = new Bundle {
val in = node.bundleIn
val out = node.bundleOut
}
((io.in zip io.out) zip (node.edgesIn zip node.edgesOut)) foreach { case ((in, out), (edgeIn, edgeOut)) =>
val beatBytes = edgeOut.manager.beatBytes
val d_send = RegInit(Bool(false))
val d_recv = RegInit(Bool(false))
val d_pause = RegInit(Bool(true))
val d_error = RegInit(Bool(false))
val d_write = RegInit(Bool(false))
val d_addr = Reg(in.haddr)
val d_size = Reg(in.hsize)
when (out.d.valid) { d_recv := Bool(false) }
when (out.a.ready) { d_send := Bool(false) }
when (in.hresp) { d_pause := Bool(false) }
val a_count = RegInit(UInt(0, width = 4))
val a_first = a_count === UInt(0)
val d_last = a_first
val burst_sizes = Seq(1, 1, 4, 4, 8, 8, 16, 16)
val a_burst_size = Vec(burst_sizes.map(beats => UInt(log2Ceil(beats * beatBytes))))(in.hburst)
val a_burst_mask = Vec(burst_sizes.map(beats => UInt(beats * beatBytes - 1)))(in.hburst)
val a_burst_ok =
in.htrans === AHBParameters.TRANS_NONSEQ && // only start burst on first AHB beat
in.hsize === UInt(log2Ceil(beatBytes)) && // not a narrow burst
(in.haddr & a_burst_mask) === UInt(0) && // address aligned to burst size
in.hburst =/= AHBParameters.BURST_INCR && // we know the burst length a priori
Mux(in.hwrite, // target device supports the burst
edgeOut.manager.supportsPutFullSafe(in.haddr, a_burst_size),
edgeOut.manager.supportsGetSafe (in.haddr, a_burst_size))
val beat = TransferSizes(1, beatBytes)
val a_legal = // Is the single-beat access allowed?
Mux(in.hwrite,
edgeOut.manager.supportsPutFullSafe(in.haddr, in.hsize, Some(beat)),
edgeOut.manager.supportsGetSafe (in.haddr, in.hsize, Some(beat)))
val a_access = in.htrans === AHBParameters.TRANS_NONSEQ || in.htrans === AHBParameters.TRANS_SEQ
val a_accept = in.hready && in.hsel && a_access
when (a_accept) {
a_count := a_count - UInt(1)
when ( in.hwrite) { d_send := Bool(true) }
when (!in.hwrite) { d_recv := Bool(true) }
when (a_first) {
a_count := Mux(a_burst_ok, a_burst_mask >> log2Ceil(beatBytes), UInt(0))
d_send := a_legal
d_recv := a_legal
d_pause := Bool(true)
d_write := in.hwrite
d_addr := in.haddr
d_size := Mux(a_burst_ok, a_burst_size, in.hsize)
}
}
out.a.valid := d_send
out.a.bits.opcode := Mux(d_write, TLMessages.PutFullData, TLMessages.Get)
out.a.bits.param := UInt(0)
out.a.bits.size := d_size
out.a.bits.source := UInt(0)
out.a.bits.address := d_addr
out.a.bits.data := in.hwdata
out.a.bits.mask := MaskGen(d_addr, d_size, beatBytes)
out.d.ready := d_recv // backpressure AccessAckData arriving faster than AHB beats
in.hrdata := out.d.bits.data
// In a perfect world, we'd use these signals
val hresp = d_error || (out.d.valid && out.d.bits.error)
val hreadyout = Mux(d_write, (!d_send || out.a.ready) && (!d_last || !d_recv || out.d.valid), out.d.valid || !d_recv)
// Make the error persistent (and defer it to the last beat--otherwise AHB can cancel the burst!)
d_error :=
(hresp && !(a_first && in.hready)) || // clear error when a new beat starts
(a_accept && !a_legal) // error if the address requested is illegal
// When we report an error, we need to be hreadyout LOW for one cycle
in.hresp := hreadyout && (hresp && d_last)
in.hreadyout := hreadyout && !(hresp && d_last && d_pause)
// Unused channels
out.b.ready := Bool(true)
out.c.valid := Bool(false)
out.e.valid := Bool(false)
}
}
}
object AHBToTL
{
def apply()(x: AHBOutwardNode)(implicit p: Parameters, sourceInfo: SourceInfo): TLOutwardNode = {
val tl = LazyModule(new AHBToTL)
tl.node := x
tl.node
}
}

View File

@ -0,0 +1,50 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba.ahb
import Chisel._
import freechips.rocketchip.config.Parameters
import freechips.rocketchip.diplomacy._
import freechips.rocketchip.regmapper._
import scala.math.{min,max}
class AHBFanout()(implicit p: Parameters) extends LazyModule {
val node = AHBNexusNode(
numSlavePorts = 1 to 1,
numMasterPorts = 1 to 32,
masterFn = { case Seq(m) => m },
slaveFn = { seq => seq(0).copy(slaves = seq.flatMap(_.slaves)) })
lazy val module = new LazyModuleImp(this) {
val io = new Bundle {
val in = node.bundleIn
val out = node.bundleOut
}
// Require consistent bus widths
val port0 = node.edgesOut(0).slave
node.edgesOut.foreach { edge =>
val port = edge.slave
require (port.beatBytes == port0.beatBytes,
s"${port.slaves.map(_.name)} ${port.beatBytes} vs ${port0.slaves.map(_.name)} ${port0.beatBytes}")
}
val port_addrs = node.edgesOut.map(_.slave.slaves.map(_.address).flatten)
val routingMask = AddressDecoder(port_addrs)
val route_addrs = port_addrs.map(_.map(_.widen(~routingMask)).distinct)
val in = io.in(0)
val a_sel = Vec(route_addrs.map(seq => seq.map(_.contains(in.haddr)).reduce(_ || _)))
val d_sel = Reg(a_sel)
when (in.hready) { d_sel := a_sel }
(a_sel zip io.out) foreach { case (sel, out) =>
out := in
out.hsel := in.hsel && sel
}
in.hreadyout := !Mux1H(d_sel, io.out.map(!_.hreadyout))
in.hresp := Mux1H(d_sel, io.out.map(_.hresp))
in.hrdata := Mux1H(d_sel, io.out.map(_.hrdata))
}
}

View File

@ -0,0 +1,11 @@
// See LICENSE.SiFive for license details.
package freechips.rocketchip.amba
import Chisel._
import freechips.rocketchip.diplomacy.OutwardNodeHandle
package object ahb
{
type AHBOutwardNode = OutwardNodeHandle[AHBMasterPortParameters, AHBSlavePortParameters, AHBBundle]
}