Skip to content

Commit

Permalink
Merge pull request #13652 from dotty-staging/add-quotes-change-owner
Browse files Browse the repository at this point in the history
Allow nested Quotes with a different owners
  • Loading branch information
nicolasstucki authored Mar 23, 2022
2 parents 981a34e + e0989a0 commit f88d996
Show file tree
Hide file tree
Showing 10 changed files with 190 additions and 4 deletions.
6 changes: 6 additions & 0 deletions compiler/src/scala/quoted/runtime/impl/QuotesImpl.scala
Original file line number Diff line number Diff line change
Expand Up @@ -2627,6 +2627,8 @@ class QuotesImpl private (using val ctx: Context) extends Quotes, QuoteUnpickler

def show(using printer: Printer[Symbol]): String = printer.show(self)

def asQuotes: Nested = new QuotesImpl(using ctx.withOwner(self))

end extension

private def appliedTypeRef(sym: Symbol): TypeRepr =
Expand Down Expand Up @@ -2918,6 +2920,10 @@ class QuotesImpl private (using val ctx: Context) extends Quotes, QuoteUnpickler
|which has the AST representation
|${Printer.TreeStructure.show(tree)}
|
|
|
|Tip: The owner of a tree can be changed using method `Tree.changeOwner`.
|Tip: The default owner of definitions created in quotes can be changed using method `Symbol.asQuotes`.
|""".stripMargin)
case _ => traverseChildren(t)
}.traverse(tree)
Expand Down
95 changes: 92 additions & 3 deletions library/src/scala/quoted/Quotes.scala
Original file line number Diff line number Diff line change
Expand Up @@ -525,6 +525,8 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
* `Some` containing the implementation of the method. Returns `None` the method has no implementation.
* Any definition directly inside the implementation should have `symbol` as owner.
*
* Use `Symbol.asQuotes` to create the rhs using quoted code.
*
* See also: `Tree.changeOwner`
*/
def apply(symbol: Symbol, rhsFn: List[List[Tree]] => Option[Term]): DefDef
Expand Down Expand Up @@ -602,20 +604,53 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
* Returns `None` the method has no implementation.
* Any definition directly inside the implementation should have `symbol` as owner.
*
* Use `Symbol.asQuotes` to create the rhs using quoted code.
*
* See also: `Tree.changeOwner`
*/
def apply(symbol: Symbol, rhs: Option[Term]): ValDef
def copy(original: Tree)(name: String, tpt: TypeTree, rhs: Option[Term]): ValDef
def unapply(vdef: ValDef): (String, TypeTree, Option[Term])

/** Creates a block `{ val <name> = <rhs: Term>; <body(x): Term> }` */
/** Creates a block `{ val <name> = <rhs: Term>; <body(x): Term> }`
*
* Usage:
* ```
* ValDef.let(owner, "x", rhs1) { x =>
* ValDef.let(x.symbol.owner, "y", rhs2) { y =>
* // use `x` and `y`
* }
* }
* ```
* @syntax markdown
*/
def let(owner: Symbol, name: String, rhs: Term)(body: Ref => Term): Term

/** Creates a block `{ val x = <rhs: Term>; <body(x): Term> }` */
/** Creates a block `{ val x = <rhs: Term>; <body(x): Term> }`
*
* Usage:
* ```
* ValDef.let(owner, rhs1) { x =>
* ValDef.let(owner, rhs2) { y =>
* // use `x` and `y`
* }
* }
* ```
* @syntax markdown
*/
def let(owner: Symbol, rhs: Term)(body: Ref => Term): Term =
let(owner, "x", rhs)(body)

/** Creates a block `{ val x1 = <terms(0): Term>; ...; val xn = <terms(n-1): Term>; <body(List(x1, ..., xn)): Term> }` */
/** Creates a block `{ val x1 = <terms(0): Term>; ...; val xn = <terms(n-1): Term>; <body(List(x1, ..., xn)): Term> }`
*
* Usage:
* ```
* ValDef.let(owner, rhsList) { xs =>
* ...
* }
* ```
* @syntax markdown
*/
def let(owner: Symbol, terms: List[Term])(body: List[Ref] => Term): Term
}

Expand Down Expand Up @@ -1341,6 +1376,28 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
* ```scala sc:nocompile
* Block((DefDef(_, _, params :: Nil, _, Some(rhsFn(meth, paramRefs)))) :: Nil, Closure(meth, _))
* ```
*
* Usage:
* ```
* val mtpe = MethodType(List("arg1"))(_ => List(TypeRepr.of[Int]), _ => TypeRepr.of[Int])
* Lambda(owner, mtpe, {
* case (methSym, List(arg1: Term)) =>
* ValDef.let(methSym, f(arg1)) { ... }
* }
* )
* ```
*
* Usage with quotes:
* ```
* val mtpe = MethodType(List("arg1"))(_ => List(TypeRepr.of[Int]), _ => TypeRepr.of[Int])
* Lambda(owner, mtpe, {
* case (methSym, List(arg1: Term)) =>
* given Quotes = methSym.asQuotes
* '{ ... }
* }
* )
* ```
*
* @param owner: owner of the generated `meth` symbol
* @param tpe: Type of the definition
* @param rhsFn: Function that receives the `meth` symbol and the a list of references to the `params`
Expand Down Expand Up @@ -3817,6 +3874,35 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>

/** Case class or case object children of a sealed trait or cases of an `enum`. */
def children: List[Symbol]

/** Returns a nested quote with this symbol as splice owner (`Symbol.spliceOwner`).
*
* Changes the owner under which the definition in a quote are created.
*
* Usages:
* ```scala
* def rhsExpr(using Quotes): Expr[Unit] = '{ val y = ???; (y, y) }
* def aValDef(using Quotes)(owner: Symbol) =
* val sym = Symbol.newVal(owner, "x", TypeRepr.of[Unit], Flags.EmptyFlags, Symbol.noSymbol)
* val rhs = rhsExpr(using sym.asQuotes).asTerm
* ValDef(sym, Some(rhs))
* ```
*
* ```scala
* new TreeMap:
* override def transformTerm(tree: Term)(owner: Symbol): Term =
* tree match
* case tree: Ident =>
* given Quotes = owner.asQuotes
* // Definitions contained in the quote will be owned by `owner`.
* // No need to use `changeOwner` in this case.
* '{ val x = ???; x }.asTerm
* ```
* @syntax markdown
*/
@experimental
def asQuotes: Nested

end extension
}

Expand Down Expand Up @@ -4497,6 +4583,9 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
* override def transformTree(tree: Tree)(owner: Symbol): Tree = ???
* }
* ```
*
* Use `Symbol.asQuotes` to create quotes with the correct owner within the TreeMap.
*
* @syntax markdown
*/
trait TreeMap:
Expand Down
2 changes: 2 additions & 0 deletions project/MiMaFilters.scala
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,8 @@ object MiMaFilters {
val Library: Seq[ProblemFilter] = Seq(
// Experimental APIs that can be added in 3.2.0
ProblemFilters.exclude[DirectMissingMethodProblem]("scala.runtime.Tuples.append"),
ProblemFilters.exclude[ReversedMissingMethodProblem]("scala.quoted.Quotes#reflectModule#SymbolMethods.asQuotes"),
ProblemFilters.exclude[DirectMissingMethodProblem]("scala.quoted.Quotes#reflectModule#SymbolMethods.asQuotes"),
ProblemFilters.exclude[ReversedMissingMethodProblem]("scala.quoted.Quotes#reflectModule#TypeReprMethods.substituteTypes"),
ProblemFilters.exclude[DirectMissingMethodProblem]("scala.quoted.Quotes#reflectModule#TypeReprMethods.substituteTypes"),
ProblemFilters.exclude[ReversedMissingMethodProblem]("scala.quoted.Quotes#reflectModule#TypeReprMethods.typeArgs"),
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -85,7 +85,7 @@ class SnippetCompiler(

private def additionalMessages(wrappedSnippet: WrappedSnippet, arg: SnippetCompilerArg, sourceFile: SourceFile, context: Context): Seq[SnippetCompilerMessage] = {
Option.when(arg.flag == SCFlags.Fail && !context.reporter.hasErrors)(
SnippetCompilerMessage(None, "Snippet should not compile but compiled succesfully", MessageLevel.Error)
SnippetCompilerMessage(None, "Snippet should not compile but compiled successfully", MessageLevel.Error)
).toList
}

Expand Down
20 changes: 20 additions & 0 deletions tests/pos-macros/i13571/Macro_1.scala
Original file line number Diff line number Diff line change
@@ -0,0 +1,20 @@
import scala.quoted.*

inline def checked2[A](inline n: A): A =
${ checkedImpl2[A]('{n}) }

private def checkedImpl2[A](n: Expr[A])(using Quotes, Type[A]): Expr[A] =
import quotes.reflect.*
val tree: Term = n.asTerm
val acc = new TreeMap:
override def transformTerm(tree: Term)(owner: Symbol): Term =
tree match
case Apply(Select(x, "*"), List(y)) =>
given Quotes = owner.asQuotes
'{
val xt = ${x.asExprOf[Long]}
xt
}.asTerm
case _ =>
super.transformTerm(tree)(owner)
acc.transformTerm(tree)(Symbol.spliceOwner).asExprOf[A]
6 changes: 6 additions & 0 deletions tests/pos-macros/i13571/Test_2.scala
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
def test = {
val u = 3L
checked2(List(1L, 2L).map { k =>
u * 2L
})
}
22 changes: 22 additions & 0 deletions tests/pos-macros/i13571b/Macro_1.scala
Original file line number Diff line number Diff line change
@@ -0,0 +1,22 @@
import scala.quoted.*

inline def checked2[A](inline n: A): A =
${ checkedImpl2[A]('{n}) }

private def checkedImpl2[A](n: Expr[A])(using Quotes, Type[A]): Expr[A] =
import quotes.reflect.*
val tree: Term = n.asTerm
val acc = new TreeMap:
override def transformTerm(tree: Term)(owner: Symbol): Term =
tree match
case Apply(Select(x, "*"), List(y)) =>
bindLong(x.asExprOf[Long])(using owner.asQuotes).asTerm
case _ =>
super.transformTerm(tree)(owner)
acc.transformTerm(tree)(Symbol.spliceOwner).asExprOf[A]

def bindLong(expr: Expr[Long])(using Quotes): Expr[Long] =
'{
val xt = $expr
xt
}
6 changes: 6 additions & 0 deletions tests/pos-macros/i13571b/Test_2.scala
Original file line number Diff line number Diff line change
@@ -0,0 +1,6 @@
def test = {
val u = 3L
checked2(List(1L, 2L).map { k =>
u * 2L
})
}
34 changes: 34 additions & 0 deletions tests/pos-macros/i13922/Macro_1.scala
Original file line number Diff line number Diff line change
@@ -0,0 +1,34 @@
import scala.quoted.*

inline def optPrettyPrinter[T]: Option[T] => Option[T] =
${ optPrettyPrinterImpl[T] }

private def optPrettyPrinterImpl[T: Type](using Quotes): Expr[Option[T] => Option[T]] = {
import quotes.reflect.*

val tpe = TypeRepr.of[T]

val fn = Lambda(
Symbol.spliceOwner,
MethodType(List("macroVal"))(
_ => List(tpe),
_ => tpe
),
{
case (m, List(arg: Term)) =>
given Quotes = m.asQuotes
ValDef.let(m, "v", arg) { v =>
'{
val vv = ${ v.asExprOf[T] }
println("v=" + vv.toString())
vv
}.asTerm
}

case _ =>
report.errorAndAbort("Fails compile")
}
).asExprOf[T => T]

'{ (_: Option[T]).map(${ fn }) }
}
1 change: 1 addition & 0 deletions tests/pos-macros/i13922/Test_2.scala
Original file line number Diff line number Diff line change
@@ -0,0 +1 @@
def test = optPrettyPrinter(Some("foo"))

0 comments on commit f88d996

Please sign in to comment.