package shapeless
package ops
import scala.annotation.tailrec
import poly._
object hlist {
trait IsHCons[L <: HList] {
type H
type T <: HList
def head(l : L) : H
def tail(l : L) : T
}
object IsHCons {
type Aux[L <: HList, H0, T0 <: HList] = IsHCons[L] { type H = H0; type T = T0 }
implicit def hlistIsHCons[H0, T0 <: HList]: Aux[H0 :: T0, H0, T0] =
new IsHCons[H0 :: T0] {
type H = H0
type T = T0
def head(l : H0 :: T0) : H = l.head
def tail(l : H0 :: T0) : T = l.tail
}
}
trait Mapped[L <: HList, F[_]] {
type Out <: HList
}
object Mapped {
type Aux[L <: HList, F[_], Out0 <: HList] = Mapped[L, F] { type Out = Out0 }
implicit def hnilMapped[F[_]]: Aux[HNil, F, HNil] = new Mapped[HNil, F] { type Out = HNil }
implicit def hlistIdMapped[L <: HList]: Aux[L, Id, L] = new Mapped[L, Id] { type Out = L }
implicit def hlistMapped1[H, T <: HList, F[_], OutM <: HList](implicit mt : Mapped.Aux[T, F, OutM]): Aux[H :: T, F, F[H] :: OutM] =
new Mapped[H :: T, F] { type Out = F[H] :: OutM }
implicit def hlistMapped2[H, T <: HList, F, OutM <: HList](implicit mt : Mapped.Aux[T, Const[F]#λ, OutM]): Aux[H :: T, Const[F]#λ, F :: OutM] =
new Mapped[H :: T, Const[F]#λ] { type Out = F :: OutM }
}
trait Comapped[L <: HList, F[_]] {
type Out <: HList
}
trait LowPriorityComapped {
type Aux[L <: HList, F[_], Out0 <: HList] = Comapped[L, F] { type Out = Out0 }
implicit def hlistIdComapped[L <: HList]: Aux[L, Id, L] = new Comapped[L, Id] { type Out = L }
}
object Comapped extends LowPriorityComapped {
implicit def hnilComapped[F[_]]: Aux[HNil, F, HNil] = new Comapped[HNil, F] { type Out = HNil }
implicit def hlistComapped[H, T <: HList, F[_]](implicit mt : Comapped[T, F]): Aux[F[H] :: T, F, H :: mt.Out] =
new Comapped[F[H] :: T, F] { type Out = H :: mt.Out }
}
trait NatTRel[L1 <: HList, F1[_], L2 <: HList, F2[_]] {
def map(nt: F1 ~> F2, fa: L1): L2
}
object NatTRel {
implicit def hnilNatTRel1[F1[_], F2[_]] = new NatTRel[HNil, F1, HNil, F2] {
def map(f: F1 ~> F2, fa: HNil): HNil = HNil
}
implicit def hnilNatTRel2[F1[_], H2] = new NatTRel[HNil, F1, HNil, Const[H2]#λ] {
def map(f: F1 ~> Const[H2]#λ, fa: HNil): HNil = HNil
}
implicit def hlistNatTRel1[H, F1[_], F2[_], T1 <: HList, T2 <: HList](implicit nt : NatTRel[T1, F1, T2, F2]) =
new NatTRel[F1[H] :: T1, F1, F2[H] :: T2, F2] {
def map(f: F1 ~> F2, fa: F1[H] :: T1): F2[H] :: T2 = f(fa.head) :: nt.map(f, fa.tail)
}
implicit def hlistNatTRel2[H, F2[_], T1 <: HList, T2 <: HList](implicit nt : NatTRel[T1, Id, T2, F2]) =
new NatTRel[H :: T1, Id, F2[H] :: T2, F2] {
def map(f: Id ~> F2, fa: H :: T1): F2[H] :: T2 = f(fa.head) :: nt.map(f, fa.tail)
}
implicit def hlistNatTRel3[H, F1[_], T1 <: HList, T2 <: HList](implicit nt : NatTRel[T1, F1, T2, Id]) =
new NatTRel[F1[H] :: T1, F1, H :: T2, Id] {
def map(f: F1 ~> Id, fa: F1[H] :: T1): H :: T2 = f(fa.head) :: nt.map(f, fa.tail)
}
implicit def hlistNatTRel4[H1, F1[_], T1 <: HList, H2, T2 <: HList](implicit nt : NatTRel[T1, F1, T2, Const[H2]#λ]) =
new NatTRel[F1[H1] :: T1, F1, H2 :: T2, Const[H2]#λ] {
def map(f: F1 ~> Const[H2]#λ, fa: F1[H1] :: T1): H2 :: T2 = f(fa.head) :: nt.map(f, fa.tail)
}
implicit def hlistNatTRel5[H1, T1 <: HList, H2, T2 <: HList](implicit nt : NatTRel[T1, Id, T2, Const[H2]#λ]) =
new NatTRel[H1 :: T1, Id, H2 :: T2, Const[H2]#λ] {
def map(f: Id ~> Const[H2]#λ, fa: H1 :: T1): H2 :: T2 = f(fa.head) :: nt.map(f, fa.tail)
}
}
trait HKernel {
type L <: HList
type Mapped[G[_]] <: HList
type Length <: Nat
def map[F[_], G[_]](f: F ~> G, l: Mapped[F]): Mapped[G]
def tabulate[C](from: Int)(f: Int => C): Mapped[Const[C]#λ]
def toList[C](l: Mapped[Const[C]#λ]): List[C]
def length: Int
}
trait HNilHKernel extends HKernel {
type L = HNil
type Mapped[G[_]] = HNil
type Length = _0
def map[F[_], G[_]](f: F ~> G, l: HNil): HNil = HNil
def tabulate[C](from: Int)(f: Int => C): HNil = HNil
def toList[C](l: HNil): List[C] = Nil
def length: Int = 0
}
case object HNilHKernel extends HNilHKernel
final case class HConsHKernel[H, T <: HKernel](tail: T) extends HKernel {
type L = H :: tail.L
type Mapped[G[_]] = G[H] :: tail.Mapped[G]
type Length = Succ[tail.Length]
def map[F[_], G[_]](f: F ~> G, l: F[H] :: tail.Mapped[F]): G[H] :: tail.Mapped[G] = f(l.head) :: tail.map(f, l.tail)
def tabulate[C](from: Int)(f: Int => C): C :: tail.Mapped[Const[C]#λ] = f(from) :: tail.tabulate(from+1)(f)
def toList[C](l: C :: tail.Mapped[Const[C]#λ]): List[C] = l.head :: tail.toList(l.tail)
def length: Int = 1+tail.length
}
object HKernel {
def apply[L <: HList](implicit mk: HKernelAux[L]): mk.Out = mk()
def apply[L <: HList](l: L)(implicit mk: HKernelAux[L]): mk.Out = mk()
}
trait HKernelAux[L <: HList] {
type Out <: HKernel
def apply(): Out
}
object HKernelAux {
implicit def mkHNilHKernel = new HKernelAux[HNil] {
type Out = HNilHKernel
def apply() = HNilHKernel
}
implicit def mkHListHKernel[H, T <: HList](implicit ct: HKernelAux[T]) = new HKernelAux[H :: T] {
type Out = HConsHKernel[H, ct.Out]
def apply() = HConsHKernel[H, ct.Out](ct())
}
}
trait Length[L <: HList] extends DepFn0 { type Out <: Nat }
object Length {
import shapeless.nat._
type Aux[L <: HList, N <: Nat] = Length[L] { type Out = N }
implicit def hnilLength: Aux[HNil, _0] = new Length[HNil] {
type Out = _0
def apply(): Out = _0
}
implicit def hlistLength[H, T <: HList, N <: Nat](implicit lt : Aux[T, N], sn : Witness.Aux[Succ[N]]): Aux[H :: T, Succ[N]] = new Length[H :: T] {
type Out = Succ[N]
def apply(): Out = sn.value
}
}
trait Mapper[HF, In <: HList] extends DepFn1[In] { type Out <: HList }
object Mapper {
type Aux[HF, In <: HList, Out0 <: HList] = Mapper[HF, In] { type Out = Out0 }
implicit def hnilMapper1[HF]: Aux[HF, HNil, HNil] =
new Mapper[HF, HNil] {
type Out = HNil
def apply(l : HNil): Out = HNil
}
implicit def hlistMapper1[HF <: Poly, InH, InT <: HList]
(implicit hc : Case1[HF, InH], mt : Mapper[HF, InT]): Aux[HF, InH :: InT, hc.Result :: mt.Out] =
new Mapper[HF, InH :: InT] {
type Out = hc.Result :: mt.Out
def apply(l : InH :: InT): Out = hc(l.head) :: mt(l.tail)
}
}
trait FlatMapper[HF, In <: HList] extends DepFn1[In] { type Out <: HList }
object FlatMapper {
type Aux[HF, In <: HList, Out0 <: HList] = FlatMapper[HF, In] { type Out = Out0 }
implicit def hnilFlatMapper1[HF]: Aux[HF, HNil, HNil] =
new FlatMapper[HF, HNil] {
type Out = HNil
def apply(l : HNil): Out = HNil
}
implicit def hlistFlatMapper1[HF <: Poly, InH, OutH <: HList, InT <: HList, OutT <: HList, Out0 <: HList]
(implicit
hc : Case1.Aux[HF, InH, OutH],
mt : FlatMapper.Aux[HF, InT, OutT],
prepend : Prepend.Aux[OutH, OutT, Out0]
): Aux[HF, InH :: InT, Out0] =
new FlatMapper[HF, InH :: InT] {
type Out = Out0
def apply(l : InH :: InT): Out = prepend(hc(l.head), mt(l.tail))
}
}
trait ConstMapper[C, L <: HList] extends DepFn2[C, L] { type Out <: HList }
object ConstMapper {
type Aux[C, L <: HList, Out0 <: HList] = ConstMapper[C, L] { type Out = Out0 }
implicit def hnilConstMapper[C]: Aux[C, HNil, HNil] =
new ConstMapper[C, HNil] {
type Out = HNil
def apply(c : C, l : HNil): Out = l
}
implicit def hlistConstMapper[H, T <: HList, C]
(implicit mct : ConstMapper[C, T]): Aux[C, H :: T, C :: mct.Out] =
new ConstMapper[C, H :: T] {
type Out = C :: mct.Out
def apply(c : C, l : H :: T): Out = c :: mct(c, l.tail)
}
}
trait MapFolder[L <: HList, R, HF] {
def apply(l : L, in : R, op : (R, R) => R) : R
}
object MapFolder {
implicit def hnilMapFolder[R, HF]: MapFolder[HNil, R, HF] = new MapFolder[HNil, R, HF] {
def apply(l : HNil, in : R, op : (R, R) => R): R = in
}
implicit def hlistMapFolder[H, T <: HList, R, HF <: Poly]
(implicit hc : Case1.Aux[HF, H, R], tf : MapFolder[T, R, HF]): MapFolder[H :: T, R, HF] =
new MapFolder[H :: T, R, HF] {
def apply(l : H :: T, in : R, op : (R, R) => R): R = op(hc(l.head), tf(l.tail, in, op))
}
}
trait LeftFolder[L <: HList, In, HF] extends DepFn2[L, In]
object LeftFolder {
type Aux[L <: HList, In, HF, Out0] = LeftFolder[L, In, HF] { type Out = Out0 }
implicit def hnilLeftFolder[In, HF]: Aux[HNil, In , HF, In] =
new LeftFolder[HNil, In, HF] {
type Out = In
def apply(l : HNil, in : In): Out = in
}
implicit def hlistLeftFolder[H, T <: HList, In, HF, OutH]
(implicit f : Case2.Aux[HF, In, H, OutH], ft : LeftFolder[T, OutH, HF]): Aux[H :: T, In, HF, ft.Out] =
new LeftFolder[H :: T, In, HF] {
type Out = ft.Out
def apply(l : H :: T, in : In) : Out = ft(l.tail, f(in, l.head))
}
}
trait RightFolder[L <: HList, In, HF] extends DepFn2[L, In]
object RightFolder {
type Aux[L <: HList, In, HF, Out0] = RightFolder[L, In, HF] { type Out = Out0 }
implicit def hnilRightFolder[In, HF]: Aux[HNil, In, HF, In] =
new RightFolder[HNil, In, HF] {
type Out = In
def apply(l : HNil, in : In): Out = in
}
implicit def hlistRightFolder[H, T <: HList, In, HF, OutT]
(implicit ft : RightFolder.Aux[T, In, HF, OutT], f : Case2[HF, H, OutT]): Aux[H :: T, In, HF, f.Result] =
new RightFolder[H :: T, In, HF] {
type Out = f.Result
def apply(l : H :: T, in : In): Out = f(l.head, ft(l.tail, in))
}
}
trait LeftReducer[L <: HList, HF] extends DepFn1[L]
object LeftReducer {
type Aux[L <: HList, HF, Out0] = LeftReducer[L, HF] { type Out = Out0 }
implicit def leftReducer[H, T <: HList, HF](implicit folder : LeftFolder[T, H, HF]): Aux[H :: T, HF, folder.Out] =
new LeftReducer[H :: T, HF] {
type Out = folder.Out
def apply(l : H :: T) : Out = folder.apply(l.tail, l.head)
}
}
trait RightReducer[L <: HList, HF] extends DepFn1[L]
object RightReducer {
type Aux[L <: HList, HF, Out0] = RightReducer[L, HF] { type Out = Out0 }
implicit def hsingleRightReducer[H, HF]: Aux[H :: HNil, HF, H] =
new RightReducer[H :: HNil, HF] {
type Out = H
def apply(l : H :: HNil): Out = l.head
}
implicit def hlistRightReducer[H, T <: HList, HF, OutT]
(implicit rt : RightReducer.Aux[T, HF, OutT], f : Case2[HF, H, OutT]): Aux[H :: T, HF, f.Result] =
new RightReducer[H :: T, HF] {
type Out = f.Result
def apply(l : H :: T): Out = f(l.head, rt(l.tail))
}
}
trait Unifier[L <: HList] extends DepFn1[L] { type Out <: HList }
object Unifier {
type Aux[L <: HList, Out0 <: HList] = Unifier[L] { type Out = Out0 }
implicit val hnilUnifier: Aux[HNil, HNil] = new Unifier[HNil] {
type Out = HNil
def apply(l : HNil): Out = l
}
implicit def hsingleUnifier[T]: Aux[T :: HNil, T :: HNil] =
new Unifier[T :: HNil] {
type Out = T :: HNil
def apply(l : T :: HNil): Out = l
}
implicit def hlistUnifier[H1, H2, L, T <: HList]
(implicit u : Lub[H1, H2, L], lt : Unifier[L :: T]): Aux[H1 :: H2 :: T, L :: lt.Out] =
new Unifier[H1 :: H2 :: T] {
type Out = L :: lt.Out
def apply(l : H1 :: H2 :: T): Out = u.left(l.head) :: lt(u.right(l.tail.head) :: l.tail.tail)
}
}
trait SubtypeUnifier[L <: HList, B] extends DepFn1[L] { type Out <: HList }
object SubtypeUnifier {
type Aux[L <: HList, B, Out0 <: HList] = SubtypeUnifier[L, B] { type Out = Out0 }
implicit def hnilSubtypeUnifier[B]: Aux[HNil, B, HNil] =
new SubtypeUnifier[HNil, B] {
type Out = HNil
def apply(l : HNil): Out = l
}
implicit def hlistSubtypeUnifier1[H, T <: HList, B]
(implicit st : H <:< B, sut: SubtypeUnifier[T, B]): Aux[H :: T, B, B :: sut.Out] =
new SubtypeUnifier[H :: T, B] {
type Out = B :: sut.Out
def apply(l : H :: T): Out = st(l.head) :: sut(l.tail)
}
implicit def hlistSubtypeUnifier2[H, T <: HList, B]
(implicit nst : H <:!< B, sut: SubtypeUnifier[T, B]): Aux[H :: T, B, H :: sut.Out] =
new SubtypeUnifier[H :: T, B] {
type Out = H :: sut.Out
def apply(l : H :: T): Out = l.head :: sut(l.tail)
}
}
trait ToList[-L <: HList, Lub] {
def apply(l: L): List[Lub]
}
trait LowPriorityToList {
implicit def hlistToListAny[L <: HList]: ToList[L, Any] =
new ToList[L, Any] {
type Out = List[Any]
val b = scala.collection.mutable.ListBuffer.empty[Any]
def apply(l : L): Out = {
@tailrec
def loop(l : HList): Unit = l match {
case hd :: tl => b += hd ; loop(tl)
case _ =>
}
loop(l)
b.toList
}
}
}
object ToList extends LowPriorityToList {
implicit def hnilToList[T] : ToList[HNil, T] =
new ToList[HNil, T] {
type Out = List[T]
def apply(l : HNil): Out = Nil
}
implicit def hsingleToList[T] : ToList[T :: HNil, T] =
new ToList[T :: HNil, T] {
type Out = List[T]
def apply(l : T :: HNil): Out = List(l.head)
}
implicit def hlistToList[H1, H2, T <: HList, L]
(implicit u : Lub[H1, H2, L], ttl : ToList[H2 :: T, L]): ToList[H1 :: H2 :: T, L] =
new ToList[H1 :: H2 :: T, L] {
type Out = List[L]
def apply(l : H1 :: H2 :: T): Out = u.left(l.head) :: ttl(l.tail)
}
}
trait ToArray[-L <: HList, Lub] {
def apply(len : Int, l : L, i : Int) : Array[Lub]
}
trait LowPriorityToArray {
implicit def hlistToArrayAnyRef[L <: HList]: ToArray[L, Any] =
new ToArray[L, Any] {
def apply(len: Int, l : L, i : Int) : Array[Any] = {
val arr = Array[Any](len)
@tailrec
def loop(l : HList, i: Int): Unit = l match {
case hd :: tl => arr(i) = hd ; loop(tl, i+1)
case _ =>
}
loop(l, 0)
arr
}
}
}
object ToArray extends LowPriorityToArray {
import scala.reflect.ClassTag
implicit def hnilToArray[T : ClassTag] : ToArray[HNil, T] =
new ToArray[HNil, T] {
def apply(len : Int, l : HNil, i : Int) = Array.ofDim[T](len)
}
implicit def hsingleToArray[T : ClassTag] : ToArray[T :: HNil, T] =
new ToArray[T :: HNil, T] {
def apply(len : Int, l : T :: HNil, i : Int) = {
val arr = Array.ofDim[T](len)
arr(i) = l.head
arr
}
}
implicit def hlistToArray[H1, H2, T <: HList, L]
(implicit u : Lub[H1, H2, L], tta : ToArray[H2 :: T, L]): ToArray[H1 :: H2 :: T, L] =
new ToArray[H1 :: H2 :: T, L] {
def apply(len : Int, l : H1 :: H2 :: T, i : Int) = {
val arr = tta(len, l.tail, i+1)
arr(i) = u.left(l.head)
arr
}
}
}
trait Tupler[L <: HList] extends DepFn1[L]
object Tupler extends TuplerInstances {
implicit val hnilTupler: Aux[HNil, Unit] =
new Tupler[HNil] {
type Out = Unit
def apply(l: HNil): Out = ()
}
}
trait Last[L <: HList] extends DepFn1[L]
object Last {
type Aux[L <: HList, Out0] = Last[L] { type Out = Out0 }
implicit def hsingleLast[H]: Aux[H :: HNil, H] =
new Last[H :: HNil] {
type Out = H
def apply(l : H :: HNil): Out = l.head
}
implicit def hlistLast[H, T <: HList]
(implicit lt : Last[T]): Aux[H :: T, lt.Out] =
new Last[H :: T] {
type Out = lt.Out
def apply(l : H :: T): Out = lt(l.tail)
}
}
trait Init[L <: HList] extends DepFn1[L] { type Out <: HList }
object Init {
type Aux[L <: HList, Out0 <: HList] = Init[L] { type Out = Out0 }
implicit def hsingleInit[H]: Aux[H :: HNil, HNil] =
new Init[H :: HNil] {
type Out = HNil
def apply(l : H :: HNil): Out = HNil
}
implicit def hlistInit[H, T <: HList, OutH, OutT <: HList]
(implicit it : Init[T]): Aux[H :: T, H :: it.Out] =
new Init[H :: T] {
type Out = H :: it.Out
def apply(l : H :: T): Out = l.head :: it(l.tail)
}
}
trait Selector[L <: HList, U] extends DepFn1[L] { type Out = U }
object Selector {
type Aux[L <: HList, U] = Selector[L, U]
implicit def hlistSelect1[H, T <: HList]: Aux[H :: T, H] =
new Selector[H :: T, H] {
def apply(l : H :: T) = l.head
}
implicit def hlistSelect[H, T <: HList, U]
(implicit st : Selector[T, U]): Aux[H :: T, U] =
new Selector[H :: T, U] {
def apply(l : H :: T) = st(l.tail)
}
}
trait Filter[L <: HList, U] extends DepFn1[L] { type Out <: HList }
object Filter {
type Aux[L <: HList, U, Out0 <: HList] = Filter[L, U] { type Out = Out0 }
implicit def hlistFilterHNil[L <: HList, U]: Aux[HNil, U, HNil] =
new Filter[HNil, U] {
type Out = HNil
def apply(l : HNil): Out = HNil
}
implicit def hlistFilter1[L <: HList, H]
(implicit f : Filter[L, H]): Aux[H :: L, H, H :: f.Out] =
new Filter[H :: L, H] {
type Out = H :: f.Out
def apply(l : H :: L) : Out = l.head :: f(l.tail)
}
implicit def hlistFilter2[H, L <: HList, U]
(implicit f : Filter[L, U], e : U =:!= H): Aux[H :: L, U, f.Out] =
new Filter[H :: L, U] {
type Out = f.Out
def apply(l : H :: L): Out = f(l.tail)
}
}
trait FilterNot[L <: HList, U] extends DepFn1[L] { type Out <: HList }
object FilterNot {
type Aux[L <: HList, U, Out0 <: HList] = FilterNot[L, U] { type Out = Out0 }
implicit def hlistFilterNotHNil[L <: HList, U]: Aux[HNil, U, HNil] =
new FilterNot[HNil, U] {
type Out = HNil
def apply(l : HNil): Out = HNil
}
implicit def hlistFilterNot1[L <: HList, H]
(implicit f: FilterNot[L, H]): Aux[H :: L, H, f.Out] =
new FilterNot[H :: L, H] {
type Out = f.Out
def apply(l : H :: L): Out = f(l.tail)
}
implicit def hlistFilterNot2[H, L <: HList, U, Out <: HList]
(implicit f: FilterNot[L, U], e: U =:!= H): Aux[H :: L, U, H :: f.Out] =
new FilterNot[H :: L, U] {
type Out = H :: f.Out
def apply(l : H :: L): Out = l.head :: f(l.tail)
}
}
trait Remove[L <: HList, E] extends DepFn1[L]
object Remove {
type Aux[L <: HList, E, Out0] = Remove[L, E] { type Out = Out0 }
implicit def hlistRemove1[H, T <: HList]: Aux[H :: T, H, (H, T)] =
new Remove[H :: T, H] {
type Out = (H, T)
def apply(l : H :: T): Out = (l.head, l.tail)
}
implicit def hlistRemove[H, T <: HList, E, OutT <: HList](implicit r : Aux[T, E, (E, OutT)]): Aux[H :: T, E, (E, H :: OutT)] =
new Remove[H :: T, E] {
type Out = (E, H :: OutT)
def apply(l : H :: T): Out = {
val (e, tail) = r(l.tail)
(e, l.head :: tail)
}
}
}
trait RemoveAll[L <: HList, SL <: HList] extends DepFn1[L]
object RemoveAll {
type Aux[L <: HList, SL <: HList, Out0] = RemoveAll[L, SL] { type Out = Out0 }
implicit def hlistRemoveAllNil[L <: HList]: Aux[L, HNil, (HNil, L)] =
new RemoveAll[L, HNil] {
type Out = (HNil, L)
def apply(l : L): Out = (HNil, l)
}
implicit def hlistRemoveAll[L <: HList, E, RemE <: HList, Rem <: HList, SLT <: HList]
(implicit rt : Remove.Aux[L, E, (E, RemE)], st : Aux[RemE, SLT, (SLT, Rem)]): Aux[L, E :: SLT, (E :: SLT, Rem)] =
new RemoveAll[L, E :: SLT] {
type Out = (E :: SLT, Rem)
def apply(l : L): Out = {
val (e, rem) = rt(l)
val (sl, left) = st(rem)
(e :: sl, left)
}
}
}
trait Replacer[L <: HList, U, V] extends DepFn2[L, V]
object Replacer {
type Aux[L <: HList, U, V, Out0] = Replacer[L, U, V] { type Out = Out0 }
implicit def hlistReplacer1[T <: HList, U, V]: Aux[U :: T, U, V, (U, V :: T)] =
new Replacer[U :: T, U, V] {
type Out = (U, V :: T)
def apply(l : U :: T, v : V): Out = (l.head, v :: l.tail)
}
implicit def hlistReplacer2[H, T <: HList, U, V, OutT <: HList]
(implicit ut : Aux[T, U, V, (U, OutT)]): Aux[H :: T, U, V, (U, H :: OutT)] =
new Replacer[H :: T, U, V] {
type Out = (U, H :: OutT)
def apply(l : H :: T, v : V): Out = {
val (u, outT) = ut(l.tail, v)
(u, l.head :: outT)
}
}
}
trait ReplaceAt[L <: HList, N <: Nat, V] extends DepFn2[L, V]
object ReplaceAt {
type Aux[L <: HList, N <: Nat, V, Out0] = ReplaceAt[L, N, V] { type Out = Out0 }
implicit def hlistReplaceAt1[H, T <: HList, V]: Aux[H :: T, _0, V, (H, V :: T)] =
new ReplaceAt[H :: T, _0, V] {
type Out = (H, V :: T)
def apply(l : H :: T, v : V): Out = (l.head, v :: l.tail)
}
implicit def hlistReplaceAt2[H, T <: HList, N <: Nat, U, V, Out0 <: HList]
(implicit ut : Aux[T, N, V, (U, Out0)]): Aux[H :: T, Succ[N], V, (U, H :: Out0)] =
new ReplaceAt[H :: T, Succ[N], V] {
type Out = (U, H :: Out0)
def apply(l : H :: T, v : V): Out = {
val (u, outT) = ut(l.tail, v)
(u, l.head :: outT)
}
}
}
trait At[L <: HList, N <: Nat] extends DepFn1[L]
object At {
type Aux[L <: HList, N <: Nat, Out0] = At[L, N] { type Out = Out0 }
implicit def hlistAtZero[H, T <: HList]: Aux[H :: T, _0, H] =
new At[H :: T, _0] {
type Out = H
def apply(l : H :: T): Out = l.head
}
implicit def hlistAtN[H, T <: HList, N <: Nat]
(implicit att : At[T, N]): Aux[H :: T, Succ[N], att.Out] =
new At[H :: T, Succ[N]] {
type Out = att.Out
def apply(l : H :: T) : Out = att(l.tail)
}
}
trait Drop[L <: HList, N <: Nat] extends DepFn1[L] { type Out <: HList }
object Drop {
type Aux[L <: HList, N <: Nat, Out0 <: HList] = Drop[L, N] { type Out = Out0 }
implicit def hlistDrop1[L <: HList]: Aux[L, _0, L] =
new Drop[L, _0] {
type Out = L
def apply(l : L): Out = l
}
implicit def hlistDrop2[H, T <: HList, N <: Nat]
(implicit dt : Drop[T, N]): Aux[H :: T, Succ[N], dt.Out] =
new Drop[H :: T, Succ[N]] {
type Out = dt.Out
def apply(l : H :: T): Out = dt(l.tail)
}
}
trait Take[L <: HList, N <: Nat] extends DepFn1[L] { type Out <: HList }
object Take {
type Aux[L <: HList, N <: Nat, Out0 <: HList] = Take[L, N] { type Out = Out0 }
implicit def hlistTake1[L <: HList]: Aux[L, _0, HNil] =
new Take[L, _0] {
type Out = HNil
def apply(l : L): Out = HNil
}
implicit def hlistTake2[H, T <: HList, N <: Nat, Out <: HList]
(implicit tt : Take[T, N]): Aux[H :: T, Succ[N], H :: tt.Out] =
new Take[H :: T, Succ[N]] {
type Out = H :: tt.Out
def apply(l : H :: T): Out = l.head :: tt(l.tail)
}
}
trait Split[L <: HList, N <: Nat] extends DepFn1[L]
object Split {
type Aux[L <: HList, N <: Nat, Out0] = Split[L, N] { type Out = Out0 }
implicit def split[L <: HList, N <: Nat, P <: HList, S <: HList]
(implicit split : Split0[HNil, L, N, P, S]): Aux[L, N, (P, S)] =
new Split[L, N] {
type Out = (P, S)
def apply(l : L): Out = split(HNil, l)
}
trait Split0[AccP <: HList, AccS <: HList, N <: Nat, P <: HList, S <: HList] {
def apply(accP : AccP, accS : AccS) : (P, S)
}
object Split0 {
implicit def hlistSplit1[P <: HList, S <: HList]: Split0[P, S, _0, P, S] =
new Split0[P, S, _0, P, S] {
def apply(accP : P, accS : S) : (P, S) = (accP, accS)
}
implicit def hlistSplit2[AccP <: HList, AccSH, AccST <: HList, N <: Nat, P <: HList, S <: HList]
(implicit st : Split0[AccP, AccST, N, P, S]): Split0[AccP, AccSH :: AccST, Succ[N], AccSH :: P, S] =
new Split0[AccP, AccSH :: AccST, Succ[N], AccSH :: P, S] {
def apply(accP : AccP, accS : AccSH :: AccST) : (AccSH :: P, S) =
st(accP, accS.tail) match { case (prefix, suffix) => (accS.head :: prefix, suffix) }
}
}
}
trait ReverseSplit[L <: HList, N <: Nat] extends DepFn1[L]
object ReverseSplit {
type Aux[L <: HList, N <: Nat, Out0] = ReverseSplit[L, N] { type Out = Out0 }
implicit def reverseSplit[L <: HList, N <: Nat, P <: HList, S <: HList]
(implicit split : ReverseSplit0[HNil, L, N, P, S]): Aux[L, N, (P, S)] =
new ReverseSplit[L, N] {
type Out = (P, S)
def apply(l : L): Out = split(HNil, l)
}
trait ReverseSplit0[AccP <: HList, AccS <: HList, N <: Nat, P, S] {
def apply(accP : AccP, accS : AccS): (P, S)
}
object ReverseSplit0 {
implicit def hlistReverseSplit1[P <: HList, S <: HList]: ReverseSplit0[P, S, _0, P, S] =
new ReverseSplit0[P, S, _0, P, S] {
def apply(accP : P, accS : S): (P, S) = (accP, accS)
}
implicit def hlistReverseSplit2[AccP <: HList, AccSH, AccST <: HList, N <: Nat, P, S]
(implicit st : ReverseSplit0[AccSH :: AccP, AccST, N, P, S]): ReverseSplit0[AccP, AccSH :: AccST, Succ[N], P, S] =
new ReverseSplit0[AccP, AccSH :: AccST, Succ[N], P, S] {
def apply(accP : AccP, accS : AccSH :: AccST): (P, S) = st(accS.head :: accP, accS.tail)
}
}
}
trait SplitLeft[L <: HList, U] extends DepFn1[L]
object SplitLeft {
type Aux[L <: HList, U, Out0] = SplitLeft[L, U] { type Out = Out0 }
implicit def splitLeft[L <: HList, U, P <: HList, S <: HList]
(implicit splitLeft : SplitLeft0[HNil, L, U, P, S]): Aux[L, U, (P, S)] =
new SplitLeft[L, U] {
type Out = (P, S)
def apply(l : L): Out = splitLeft(HNil, l)
}
trait SplitLeft0[AccP <: HList, AccS <: HList, U, P <: HList, S <: HList] {
def apply(accP : AccP, accS : AccS) : (P, S)
}
trait LowPrioritySplitLeft0 {
implicit def hlistSplitLeft1[AccP <: HList, AccSH, AccST <: HList, U, P <: HList, S <: HList]
(implicit slt : SplitLeft0[AccP, AccST, U, P, S]): SplitLeft0[AccP, AccSH :: AccST, U, AccSH :: P, S] =
new SplitLeft0[AccP, AccSH :: AccST, U, AccSH :: P, S] {
def apply(accP : AccP, accS : AccSH :: AccST): (AccSH :: P, S) =
slt(accP, accS.tail) match { case (prefix, suffix) => (accS.head :: prefix, suffix) }
}
}
object SplitLeft0 extends LowPrioritySplitLeft0 {
implicit def hlistSplitLeft2[P <: HList, SH, ST <: HList]: SplitLeft0[P, SH :: ST, SH, P, SH :: ST] =
new SplitLeft0[P, SH :: ST, SH, P, SH :: ST] {
def apply(accP : P, accS : SH :: ST) : (P, SH :: ST) = (accP, accS)
}
}
}
trait ReverseSplitLeft[L <: HList, U] extends DepFn1[L]
object ReverseSplitLeft {
type Aux[L <: HList, U, Out0] = ReverseSplitLeft[L, U] { type Out = Out0 }
implicit def reverseSplitLeft[L <: HList, U, P <: HList, S <: HList]
(implicit splitLeft : ReverseSplitLeft0[HNil, L, U, P, S]): Aux[L, U, (P, S)] =
new ReverseSplitLeft[L, U] {
type Out = (P, S)
def apply(l : L): Out = splitLeft(HNil, l)
}
trait ReverseSplitLeft0[AccP <: HList, AccS <: HList, U, P, S] {
def apply(accP : AccP, accS : AccS): (P, S)
}
trait LowPriorityReverseSplitLeft0 {
implicit def hlistReverseSplitLeft1[AccP <: HList, AccSH, AccST <: HList, U, P, S]
(implicit slt : ReverseSplitLeft0[AccSH :: AccP, AccST, U, P, S]): ReverseSplitLeft0[AccP, AccSH :: AccST, U, P, S] =
new ReverseSplitLeft0[AccP, AccSH :: AccST, U, P, S] {
def apply(accP : AccP, accS : AccSH :: AccST): (P, S) = slt(accS.head :: accP, accS.tail)
}
}
object ReverseSplitLeft0 extends LowPriorityReverseSplitLeft0 {
implicit def hlistReverseSplitLeft2[P <: HList, SH, ST <: HList]: ReverseSplitLeft0[P, SH :: ST, SH, P, SH :: ST] =
new ReverseSplitLeft0[P, SH :: ST, SH, P, SH :: ST] {
def apply(accP : P, accS : SH :: ST) : (P, SH :: ST) = (accP, accS)
}
}
}
trait SplitRight[L <: HList, U] extends DepFn1[L]
object SplitRight {
type Aux[L <: HList, U, Out0] = SplitRight[L, U] { type Out = Out0 }
implicit def splitRight[L <: HList, U, P <: HList, S <: HList]
(implicit splitRight : SplitRight0[L, HNil, HNil, U, P, S]): Aux[L, U, (P, S)] =
new SplitRight[L, U] {
type Out = (P, S)
def apply(l : L): Out = splitRight(l, HNil, HNil)
}
trait SplitRight0[Rev <: HList, AccP <: HList, AccS <: HList, U, P <: HList, S <: HList] {
def apply(rev : Rev, accP : AccP, accS : AccS): (P, S)
}
trait LowPrioritySplitRight0 {
implicit def hlistSplitRight1[RevH, RevT <: HList, AccP <: HList, U, P <: HList, S <: HList]
(implicit srt : SplitRight0[RevT, RevH :: AccP, HNil, U, P, S]): SplitRight0[RevH :: RevT, AccP, HNil, U, P, S] =
new SplitRight0[RevH :: RevT, AccP, HNil, U, P, S] {
def apply(rev : RevH :: RevT, accP : AccP, accS : HNil): (P, S) = srt(rev.tail, rev.head :: accP, accS)
}
implicit def hlistSplitRight2[AccPH, AccPT <: HList, AccS <: HList, U, P <: HList, S <: HList]
(implicit srt : SplitRight0[HNil, AccPT, AccPH :: AccS, U, P, S]): SplitRight0[HNil, AccPH :: AccPT, AccS, U, P, S] =
new SplitRight0[HNil, AccPH :: AccPT, AccS, U, P, S] {
def apply(rev : HNil, accP : AccPH :: AccPT, accS : AccS): (P, S) = srt(rev, accP.tail, accP.head :: accS)
}
}
object SplitRight0 extends LowPrioritySplitRight0 {
implicit def hlistSplitRight3[PH, PT <: HList, S <: HList]
(implicit reverse : Reverse[PH :: PT]): SplitRight0[HNil, PH :: PT, S, PH, reverse.Out, S] =
new SplitRight0[HNil, PH :: PT, S, PH, reverse.Out, S] {
def apply(rev : HNil, accP : PH :: PT, accS : S): (reverse.Out, S) = (accP.reverse, accS)
}
}
}
trait ReverseSplitRight[L <: HList, U] extends DepFn1[L]
object ReverseSplitRight {
type Aux[L <: HList, U, Out0] = ReverseSplitRight[L, U] { type Out = Out0 }
implicit def reverseSplitRight[L <: HList, U, P <: HList, S <: HList]
(implicit splitRight : ReverseSplitRight0[L, HNil, HNil, U, P, S]): Aux[L, U, (P, S)] =
new ReverseSplitRight[L, U] {
type Out = (P, S)
def apply(l : L): Out = splitRight(l, HNil, HNil)
}
trait ReverseSplitRight0[Rev <: HList, AccP <: HList, AccS <: HList, U, P, S] {
def apply(rev : Rev, accP : AccP, accS : AccS): (P, S)
}
trait LowPriorityReverseSplitRight0 {
implicit def hlistReverseSplitRight1[RevH, RevT <: HList, AccP <: HList, U, P <: HList, S <: HList]
(implicit srt : ReverseSplitRight0[RevT, RevH :: AccP, HNil, U, P, S]): ReverseSplitRight0[RevH :: RevT, AccP, HNil, U, P, S] =
new ReverseSplitRight0[RevH :: RevT, AccP, HNil, U, P, S] {
def apply(rev : RevH :: RevT, accP : AccP, accS : HNil): (P, S) = srt(rev.tail, rev.head :: accP, accS)
}
implicit def hlistReverseSplitRight2[AccPH, AccPT <: HList, AccS <: HList, U, P <: HList, S <: HList]
(implicit srt : ReverseSplitRight0[HNil, AccPT, AccPH :: AccS, U, P, S]): ReverseSplitRight0[HNil, AccPH :: AccPT, AccS, U, P, S] =
new ReverseSplitRight0[HNil, AccPH :: AccPT, AccS, U, P, S] {
def apply(rev : HNil, accP : AccPH :: AccPT, accS : AccS): (P, S) = srt(rev, accP.tail, accP.head :: accS)
}
}
object ReverseSplitRight0 extends LowPriorityReverseSplitRight0 {
implicit def hlistReverseSplitRight3[PH, PT <: HList, S <: HList]: ReverseSplitRight0[HNil, PH :: PT, S, PH, PH :: PT, S] =
new ReverseSplitRight0[HNil, PH :: PT, S, PH, PH :: PT, S] {
def apply(rev : HNil, accP : PH :: PT, accS : S): (PH :: PT, S) = (accP, accS)
}
}
}
trait Reverse[L <: HList] extends DepFn1[L] { type Out <: HList }
object Reverse {
type Aux[L <: HList, Out0 <: HList] = Reverse[L] { type Out = Out0 }
implicit def reverse[L <: HList, Out0 <: HList](implicit reverse : Reverse0[HNil, L, Out0]): Aux[L, Out0] =
new Reverse[L] {
type Out = Out0
def apply(l : L) : Out = reverse(HNil, l)
}
trait Reverse0[Acc <: HList, L <: HList, Out <: HList] {
def apply(acc : Acc, l : L) : Out
}
object Reverse0 {
implicit def hnilReverse[Out <: HList]: Reverse0[Out, HNil, Out] =
new Reverse0[Out, HNil, Out] {
def apply(acc : Out, l : HNil) : Out = acc
}
implicit def hlistReverse[Acc <: HList, InH, InT <: HList, Out <: HList]
(implicit rt : Reverse0[InH :: Acc, InT, Out]): Reverse0[Acc, InH :: InT, Out] =
new Reverse0[Acc, InH :: InT, Out] {
def apply(acc : Acc, l : InH :: InT) : Out = rt(l.head :: acc, l.tail)
}
}
}
trait Prepend[P <: HList, S <: HList] extends DepFn2[P, S] { type Out <: HList }
trait LowPriorityPrepend {
type Aux[P <: HList, S <: HList, Out0 <: HList] = Prepend[P, S] { type Out = Out0 }
implicit def hnilPrepend0[P <: HList, S <: HNil]: Aux[P, S, P] =
new Prepend[P, S] {
type Out = P
def apply(prefix : P, suffix : S): P = prefix
}
}
object Prepend extends LowPriorityPrepend {
implicit def hnilPrepend1[P <: HNil, S <: HList]: Aux[P, S, S] =
new Prepend[P, S] {
type Out = S
def apply(prefix : P, suffix : S): S = suffix
}
implicit def hlistPrepend[PH, PT <: HList, S <: HList]
(implicit pt : Prepend[PT, S]): Aux[PH :: PT, S, PH :: pt.Out] =
new Prepend[PH :: PT, S] {
type Out = PH :: pt.Out
def apply(prefix : PH :: PT, suffix : S): Out = prefix.head :: pt(prefix.tail, suffix)
}
}
trait ReversePrepend[P <: HList, S <: HList] extends DepFn2[P, S] { type Out <: HList }
trait LowPriorityReversePrepend {
type Aux[P <: HList, S <: HList, Out0 <: HList] = ReversePrepend[P, S] { type Out = Out0 }
implicit def hnilReversePrepend0[P <: HList, S <: HNil]
(implicit rv: Reverse[P]): Aux[P, S, rv.Out] =
new ReversePrepend[P, S] {
type Out = rv.Out
def apply(prefix: P, suffix: S) = prefix.reverse
}
}
object ReversePrepend extends LowPriorityReversePrepend {
implicit def hnilReversePrepend1[P <: HNil, S <: HList]: Aux[P, S, S] =
new ReversePrepend[P, S] {
type Out = S
def apply(prefix: P, suffix: S) = suffix
}
implicit def hlistReversePrepend[PH, PT <: HList, S <: HList]
(implicit rpt : ReversePrepend[PT, PH :: S]): Aux[PH :: PT, S, rpt.Out] =
new ReversePrepend[PH :: PT, S] {
type Out = rpt.Out
def apply(prefix : PH :: PT, suffix : S): Out = rpt(prefix.tail, prefix.head :: suffix)
}
}
trait ZipOne[H <: HList, T <: HList] extends DepFn2[H, T] { type Out <: HList }
object ZipOne {
type Aux[H <: HList, T <: HList, Out0 <: HList] = ZipOne[H, T] { type Out = Out0 }
implicit def zipOne1[H <: HList]: Aux[H, HNil, HNil] =
new ZipOne[H, HNil] {
type Out = HNil
def apply(h : H, t : HNil): Out = HNil
}
implicit def zipOne2[T <: HList]: Aux[HNil, T, HNil] =
new ZipOne[HNil, T] {
type Out = HNil
def apply(h : HNil, t : T): Out = HNil
}
implicit def zipOne3[H, T <: HList]: Aux[H :: HNil, T :: HNil, (H :: T) :: HNil] =
new ZipOne[H :: HNil, T :: HNil] {
type Out = (H :: T) :: HNil
def apply(h : H :: HNil, t : T :: HNil): Out = (h.head :: t.head) :: HNil
}
implicit def zipOne4[HH, HT <: HList, TH <: HList, TT <: HList]
(implicit zot : ZipOne[HT, TT]): Aux[HH :: HT, TH :: TT, (HH :: TH) :: zot.Out] =
new ZipOne[HH :: HT, TH :: TT] {
type Out = (HH :: TH) :: zot.Out
def apply(h : HH :: HT, t : TH :: TT): Out = (h.head :: t.head) :: zot(h.tail, t.tail)
}
}
trait Transposer[L <: HList] extends DepFn1[L] { type Out <: HList }
object Transposer {
type Aux[L <: HList, Out0 <: HList] = Transposer[L] { type Out = Out0 }
implicit def hnilTransposer: Aux[HNil, HNil] =
new Transposer[HNil] {
type Out = HNil
def apply(l : HNil): Out = l
}
implicit def hlistTransposer1[H <: HList, MC <: HList, Out0 <: HList]
(implicit mc : ConstMapper.Aux[HNil, H, MC], zo : ZipOne.Aux[H, MC, Out0]): Aux[H :: HNil, Out0] =
new Transposer[H :: HNil] {
type Out = Out0
def apply(l : H :: HNil): Out = zo(l.head, mc(HNil, l.head))
}
implicit def hlistTransposer2[H <: HList, T <: HList, OutT <: HList, Out0 <: HList]
(implicit tt : Aux[T, OutT], zo : ZipOne.Aux[H, OutT, Out0]): Aux[H :: T, Out0] =
new Transposer[H :: T] {
type Out = Out0
def apply(l : H :: T): Out = zo(l.head, tt(l.tail))
}
}
trait Zip[L <: HList] extends DepFn1[L] { type Out <: HList }
object Zip {
type Aux[L <: HList, Out0 <: HList] = Zip[L] { type Out = Out0 }
implicit def zipper[L <: HList, OutT <: HList]
(implicit
transposer : Transposer.Aux[L, OutT],
mapper : Mapper[tupled.type, OutT]): Aux[L, mapper.Out] =
new Zip[L] {
type Out = mapper.Out
def apply(l : L): Out = l.transpose map tupled
}
}
trait Unzip[L <: HList] extends DepFn1[L]
object Unzip {
type Aux[L <: HList, Out0] = Unzip[L] { type Out = Out0 }
implicit def unzipper[L <: HList, OutM <: HList, OutT <: HList]
(implicit
mapper : Mapper.Aux[productElements.type, L, OutM],
transposer : Transposer.Aux[OutM, OutT],
tupler : Tupler[OutT]): Aux[L, tupler.Out] =
new Unzip[L] {
type Out = tupler.Out
def apply(l : L): Out = (l map productElements).transpose.tupled
}
}
trait ZipApply[FL <: HList, AL <: HList] extends DepFn2[FL, AL] { type Out <: HList }
object ZipApply {
type Aux[FL <: HList, AL <: HList, Out0 <: HList] = ZipApply[FL, AL] { type Out = Out0 }
implicit def hnilZipApply: Aux[HNil, HNil, HNil] =
new ZipApply[HNil, HNil] {
type Out = HNil
def apply(fl : HNil, al : HNil): Out = HNil
}
implicit def hconsZipApply[T, R, FLT <: HList, ALT <: HList]
(implicit ztt : ZipApply[FLT, ALT]): Aux[(T => R) :: FLT, T :: ALT, R :: ztt.Out] =
new ZipApply[(T => R) :: FLT, T :: ALT] {
type Out = R :: ztt.Out
def apply(fl : (T => R) :: FLT, al : T :: ALT): Out = fl.head(al.head) :: ztt(fl.tail, al.tail)
}
}
trait ZipConst[C, L <: HList] extends DepFn2[C, L] { type Out <: HList }
object ZipConst {
type Aux[C, L <: HList, Out0 <: HList] = ZipConst[C, L] { type Out = Out0 }
implicit def constZipper[C, L <: HList, M <: HList]
(implicit
mapper: ConstMapper.Aux[C, L, M],
zipper: Zip[L :: M :: HNil]): Aux[C, L, zipper.Out] =
new ZipConst[C, L] {
type Out = zipper.Out
def apply(c: C, l: L) = zipper(l :: mapper(c, l) :: HNil)
}
}
}