Skip to content

Commit 789344e

Browse files
committed
Fix #7110: Pickle types of quote holes
To avoid wrong member selection when unpickling we must have qualifiers that have the same type as when it was pickled. This implies that we cannot expand the holes durring unplickling as the expansion of a qualifier could have a more precise type. To be able to unpickle a hole without expanding it we must know it's type, hence it must be pickled. Once the pickled quote has been unpickled, we first replace all the type splices. After the types are properly set we expand all the term holes. Changes * Types of holes are pickled in TASTy * Pickled quotes are unpickled and after that the holes are expanded * Improved debugging info for unpickled quotes
1 parent 8a42819 commit 789344e

25 files changed

+411
-60
lines changed

compiler/src/dotty/tools/dotc/ast/tpd.scala

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -42,7 +42,7 @@ object tpd extends Trees.Instance[Type] with TypedTreeInfo {
4242
Super(qual, if (mixName.isEmpty) untpd.EmptyTypeIdent else untpd.Ident(mixName), mixinClass)
4343

4444
def Apply(fn: Tree, args: List[Tree])(implicit ctx: Context): Apply = {
45-
assert(fn.isInstanceOf[RefTree] || fn.isInstanceOf[GenericApply[_]] || fn.isInstanceOf[Inlined])
45+
assert(fn.isInstanceOf[RefTree] || fn.isInstanceOf[GenericApply[_]] || fn.isInstanceOf[Inlined] || fn.isInstanceOf[tasty.TreePickler.Hole])
4646
ta.assignType(untpd.Apply(fn, args), fn, args)
4747
}
4848

compiler/src/dotty/tools/dotc/core/quoted/PickledQuotes.scala

Lines changed: 105 additions & 26 deletions
Original file line numberDiff line numberDiff line change
@@ -61,44 +61,117 @@ object PickledQuotes {
6161
}.apply(tp)
6262

6363
/** Unpickle the tree contained in the TastyExpr */
64-
def unpickleExpr(tasty: PickledQuote, args: PickledArgs)(implicit ctx: Context): Tree = {
64+
def unpickleExpr(tasty: PickledQuote, splices: PickledArgs)(implicit ctx: Context): Tree = {
6565
val tastyBytes = TastyString.unpickle(tasty)
66-
val unpickled = unpickle(tastyBytes, args, isType = false)(ctx.addMode(Mode.ReadPositions))
67-
/** Force unpickling of the tree, removes the spliced type `@quotedTypeTag type` definitions and dealiases references to `@quotedTypeTag type` */
68-
val forceAndCleanArtefacts = new TreeMap {
69-
override def transform(tree: tpd.Tree)(implicit ctx: Context): tpd.Tree = tree match {
70-
case Block(stat :: rest, expr1) if stat.symbol.hasAnnotation(defn.InternalQuoted_QuoteTypeTagAnnot) =>
71-
assert(rest.forall { case tdef: TypeDef => tdef.symbol.hasAnnotation(defn.InternalQuoted_QuoteTypeTagAnnot) })
72-
transform(expr1)
73-
case tree => super.transform(tree).withType(dealiasTypeTags(tree.tpe))
74-
}
75-
}
76-
forceAndCleanArtefacts.transform(unpickled)
66+
val unpickled = unpickle(tastyBytes, splices, isType = false)(ctx.addMode(Mode.ReadPositions))
67+
val Inlined(call, Nil, expnasion) = unpickled
68+
val inlineCtx = inlineContext(call)
69+
val expansion1 = spliceTypes(expnasion, splices)(using inlineCtx)
70+
val expansion2 = spliceTerms(expansion1, splices)(using inlineCtx)
71+
cpy.Inlined(unpickled)(call, Nil, expansion2)
7772
}
7873

7974
/** Unpickle the tree contained in the TastyType */
8075
def unpickleType(tasty: PickledQuote, args: PickledArgs)(implicit ctx: Context): Tree = {
8176
val tastyBytes = TastyString.unpickle(tasty)
8277
val unpickled = unpickle(tastyBytes, args, isType = true)(ctx.addMode(Mode.ReadPositions))
83-
val tpt = unpickled match {
84-
case Block(aliases, tpt) =>
85-
// `@quoteTypeTag type` aliases are not required after unpickling.
86-
// Type definitions are placeholders for type holes in the pickled quote, at this point
87-
// those holes have been filled. As we already dealias al references to them in `dealiasTypeTags`
88-
// there is no need to keep their definitions in the tree. As artifacts of quote reification
89-
// they also do not have a meaningful position in the source.
90-
val aliases1 = aliases.filter(!_.symbol.hasAnnotation(defn.InternalQuoted_QuoteTypeTagAnnot))
91-
seq(aliases1, tpt)
92-
case tpt => tpt
78+
spliceTypes(unpickled, args)
79+
}
80+
81+
/** Replace all term holes with the spliced terms */
82+
private def spliceTerms(tree: Tree, splices: PickledArgs)(using Context): Tree = {
83+
val evaluateHoles = new TreeMap {
84+
override def transform(tree: tpd.Tree)(implicit ctx: Context): tpd.Tree = tree match {
85+
case Hole(isTerm, idx, args) =>
86+
val reifiedArgs = args.map { arg =>
87+
if (arg.isTerm) (qctx: scala.quoted.QuoteContext) ?=> new TastyTreeExpr(arg, QuoteContext.scopeId)
88+
else new TreeType(arg, QuoteContext.scopeId)
89+
}
90+
if isTerm then
91+
val splice1 = splices(idx).asInstanceOf[Seq[Any] => scala.quoted.QuoteContext ?=> quoted.Expr[?]]
92+
val quotedExpr = splice1(reifiedArgs)(using dotty.tools.dotc.quoted.QuoteContext())
93+
val filled = PickledQuotes.quotedExprToTree(quotedExpr)
94+
95+
// We need to make sure a hole is created with the source file of the surrounding context, even if
96+
// it filled with contents a different source file.
97+
if filled.source == ctx.source then filled
98+
else filled.cloneIn(ctx.source).withSpan(tree.span)
99+
else
100+
// Replaces type holes generated by ReifyQuotes (non-spliced types).
101+
// These are types defined in a quote and used at the same level in a nested quote.
102+
val quotedType = splices(idx).asInstanceOf[Seq[Any] => quoted.Type[?]](reifiedArgs)
103+
PickledQuotes.quotedTypeToTree(quotedType)
104+
case tree: Select =>
105+
// Retain selected members
106+
val qual = transform(tree.qualifier)
107+
qual.select(tree.symbol).withSpan(tree.span)
108+
109+
case tree =>
110+
if tree.isDef then
111+
tree.symbol.annotations = tree.symbol.annotations.map {
112+
annot => annot.derivedAnnotation(transform(annot.tree))
113+
}
114+
end if
115+
116+
val tree1 = super.transform(tree)
117+
tree1.withType(mapAnnots(tree1.tpe))
118+
}
119+
120+
// Evaluate holes in type annotations
121+
private val mapAnnots = new TypeMap {
122+
override def apply(tp: Type): Type = {
123+
tp match
124+
case tp @ AnnotatedType(underlying, annot) =>
125+
val underlying1 = this(underlying)
126+
derivedAnnotatedType(tp, underlying1, annot.derivedAnnotation(transform(annot.tree)))
127+
case _ => mapOver(tp)
128+
}
129+
}
93130
}
94-
tpt.withType(dealiasTypeTags(tpt.tpe))
131+
val tree1 = evaluateHoles.transform(tree)
132+
quotePickling.println(i"**** evaluated quote\n$tree1")
133+
tree1
134+
}
135+
136+
/** Replace all type holes generated with the spliced types */
137+
private def spliceTypes(tree: Tree, splices: PickledArgs)(using Context): Tree = {
138+
tree match
139+
case Block(stat :: rest, expr1) if stat.symbol.hasAnnotation(defn.InternalQuoted_QuoteTypeTagAnnot) =>
140+
val typeSpliceMap = (stat :: rest).iterator.map {
141+
case tdef: TypeDef =>
142+
assert(tdef.symbol.hasAnnotation(defn.InternalQuoted_QuoteTypeTagAnnot))
143+
val tree = tdef.rhs match
144+
case TypeBoundsTree(_, Hole(_, idx, args), _) =>
145+
val quotedType = splices(idx).asInstanceOf[Seq[Any] => quoted.Type[?]](args)
146+
PickledQuotes.quotedTypeToTree(quotedType)
147+
case TypeBoundsTree(_, tpt, _) =>
148+
tpt
149+
(tdef.symbol, tree.tpe)
150+
}.toMap
151+
class ReplaceSplicedTyped extends TypeMap() {
152+
override def apply(tp: Type): Type = {
153+
val tp1 = tp match {
154+
case tp: TypeRef =>
155+
typeSpliceMap.get(tp.symbol) match
156+
case Some(t) if tp.typeSymbol.hasAnnotation(defn.InternalQuoted_QuoteTypeTagAnnot) => t
157+
case None => tp
158+
case _ => tp
159+
}
160+
mapOver(tp1)
161+
}
162+
}
163+
val expansion2 = new TreeTypeMap(new ReplaceSplicedTyped).transform(expr1)
164+
quotePickling.println(i"**** typed quote\n${expansion2.show}")
165+
expansion2
166+
case _ =>
167+
tree
95168
}
96169

97170
// TASTY picklingtests/pos/quoteTest.scala
98171

99172
/** Pickle tree into it's TASTY bytes s*/
100173
private def pickle(tree: Tree)(implicit ctx: Context): Array[Byte] = {
101-
quotePickling.println(i"**** pickling quote of \n${tree.show}")
174+
quotePickling.println(i"**** pickling quote of\n$tree")
102175
val pickler = new TastyPickler(defn.RootClass)
103176
val treePkl = pickler.treePkl
104177
treePkl.pickle(tree :: Nil)
@@ -109,7 +182,7 @@ object PickledQuotes {
109182
new PositionPickler(pickler, treePkl.buf.addrOfTree).picklePositions(tree :: Nil)
110183

111184
val pickled = pickler.assembleParts()
112-
quotePickling.println(s"**** pickled quote\n${new TastyPrinter(pickled).printContents()}")
185+
quotePickling.println(s"**** pickledd quote\n${new TastyPrinter(pickled).printContents()}")
113186
pickled
114187
}
115188

@@ -122,7 +195,13 @@ object PickledQuotes {
122195
unpickler.enter(Set.empty)
123196

124197
val tree = unpickler.tree
125-
quotePickling.println(i"**** unpickle quote ${tree.show}")
198+
199+
// Make sure trees and positions are fully loaded
200+
new TreeTraverser {
201+
def traverse(tree: Tree)(implicit ctx: Context): Unit = traverseChildren(tree)
202+
}.traverse(tree)
203+
204+
quotePickling.println(i"**** unpickled quote\n$tree")
126205
tree
127206
}
128207

compiler/src/dotty/tools/dotc/core/tasty/TreePickler.scala

Lines changed: 12 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -26,7 +26,8 @@ object TreePickler {
2626
override def isTerm: Boolean = isTermHole
2727
override def isType: Boolean = !isTermHole
2828
override def fallbackToText(printer: Printer): Text =
29-
s"[[$idx|" ~~ printer.toTextGlobal(args, ", ") ~~ "]]"
29+
if isTermHole then s"{{{ $idx |" ~~ printer.toTextGlobal(tpe) ~~ "|" ~~ printer.toTextGlobal(args, ", ") ~~ "}}}"
30+
else s"[[[ $idx |" ~~ printer.toTextGlobal(tpe) ~~ "|" ~~ printer.toTextGlobal(args, ", ") ~~ "]]]"
3031
}
3132
}
3233

@@ -600,9 +601,19 @@ class TreePickler(pickler: TastyPickler) {
600601
pickleTree(alias)
601602
}
602603
case Hole(_, idx, args) =>
604+
lazy val erasedSplicesType = new TypeMap() {
605+
override def apply(tp: Type): Type = tp match {
606+
case tp: TypeRef if tp.typeSymbol.isSplice || tp.typeSymbol.hasAnnotation(defn.InternalQuoted_QuoteTypeTagAnnot) => tp.dealias.typeSymbol.info.hiBound
607+
case tp =>
608+
mapOver(tp)
609+
}
610+
}
611+
val tpe = erasedSplicesType(tree.tpe)
612+
603613
writeByte(HOLE)
604614
withLength {
605615
writeNat(idx)
616+
pickleType(tree.tpe, richTypes = true)
606617
args.foreach(pickleTree)
607618
}
608619
}

compiler/src/dotty/tools/dotc/core/tasty/TreeUnpickler.scala

Lines changed: 8 additions & 31 deletions
Original file line numberDiff line numberDiff line change
@@ -1224,7 +1224,10 @@ class TreeUnpickler(reader: TastyReader,
12241224
val alias = if currentAddr == end then EmptyTree else readTpt()
12251225
TypeBoundsTree(lo, hi, alias)
12261226
case HOLE =>
1227-
readHole(end, isType = false)
1227+
val idx = readNat()
1228+
val tpe = readType()
1229+
val args = until(end)(readTerm())
1230+
TreePickler.Hole(true, idx, args).withType(tpe)
12281231
case _ =>
12291232
readPathTerm()
12301233
}
@@ -1257,7 +1260,10 @@ class TreeUnpickler(reader: TastyReader,
12571260
case HOLE =>
12581261
readByte()
12591262
val end = readEnd()
1260-
readHole(end, isType = true)
1263+
val idx = readNat()
1264+
val tpe = readType()
1265+
val args = until(end)(readTerm())
1266+
TreePickler.Hole(false, idx, args).withType(tpe)
12611267
case _ =>
12621268
if (isTypeTreeTag(nextByte)) readTerm()
12631269
else {
@@ -1299,35 +1305,6 @@ class TreeUnpickler(reader: TastyReader,
12991305
owner => new LazyReader(localReader, owner, ctx.mode, ctx.source, op)
13001306
}
13011307

1302-
def readHole(end: Addr, isType: Boolean)(implicit ctx: Context): Tree = {
1303-
val idx = readNat()
1304-
val args = until(end)(readTerm())
1305-
val splice = splices(idx)
1306-
def wrap(arg: Tree) =
1307-
if (arg.isTerm) (qctx: scala.quoted.QuoteContext) ?=> new TastyTreeExpr(arg, QuoteContext.scopeId)
1308-
else new TreeType(arg, QuoteContext.scopeId)
1309-
val reifiedArgs = args.map(wrap)
1310-
val filled = if (isType) {
1311-
val quotedType = splice.asInstanceOf[Seq[Any] => quoted.Type[?]](reifiedArgs)
1312-
PickledQuotes.quotedTypeToTree(quotedType)
1313-
}
1314-
else {
1315-
val splice1 = splice.asInstanceOf[Seq[Any] => scala.quoted.QuoteContext ?=> quoted.Expr[?]]
1316-
val quotedExpr = splice1(reifiedArgs)(using dotty.tools.dotc.quoted.QuoteContext())
1317-
PickledQuotes.quotedExprToTree(quotedExpr)
1318-
}
1319-
// We need to make sure a hole is created with the source file of the surrounding context, even if
1320-
// it filled with contents a different source file. Otherwise nodes containing holes might end
1321-
// up without a position. PositionPickler makes sure that holes always get spans assigned,
1322-
// so we can just return the filler tree with the new source and no span here.
1323-
if (filled.source == ctx.source) filled
1324-
else {
1325-
val filled1 = filled.cloneIn(ctx.source)
1326-
filled1.span = NoSpan
1327-
filled1
1328-
}
1329-
}
1330-
13311308
// ------ Setting positions ------------------------------------------------
13321309

13331310
/** Pickled span for `addr`. */

compiler/src/dotty/tools/dotc/transform/ReifyQuotes.scala

Lines changed: 29 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -331,7 +331,35 @@ class ReifyQuotes extends MacroTransform {
331331
*/
332332
private def makeHole(isTermHole: Boolean, body: Tree, splices: List[Tree], tpe: Type)(implicit ctx: Context): Hole = {
333333
val idx = embedded.addTree(body, NoSymbol)
334-
Hole(isTermHole, idx, splices).withType(tpe).asInstanceOf[Hole]
334+
335+
def getTypeHoleType(using ctx: Context) = new TypeMap() {
336+
override def apply(tp: Type): Type = tp match
337+
case tp: TypeRef if tp.typeSymbol.isSplice =>
338+
apply(tp.dealias)
339+
case tp @ TypeRef(pre, _) if pre == NoPrefix || pre.termSymbol.isLocal =>
340+
val hiBound = tp.typeSymbol.info match
341+
case info @ ClassInfo(_, _, classParents, _, _) => classParents.reduce(_ & _)
342+
case info => info.hiBound
343+
apply(hiBound)
344+
case tp =>
345+
mapOver(tp)
346+
}
347+
348+
def getTermHoleType(using ctx: Context) = new TypeMap() {
349+
override def apply(tp: Type): Type = tp match
350+
case tp @ TypeRef(NoPrefix, _) if capturers.contains(tp.symbol) =>
351+
// reference to term with a type defined in outer quote
352+
getTypeHoleType(tp)
353+
case tp @ TermRef(NoPrefix, _) if capturers.contains(tp.symbol) =>
354+
// widen term refs to terms defined in outer quote
355+
apply(tp.widenTermRefExpr)
356+
case tp =>
357+
mapOver(tp)
358+
}
359+
360+
val holeType = if isTermHole then getTermHoleType(tpe) else getTypeHoleType(tpe)
361+
362+
Hole(isTermHole, idx, splices).withType(holeType).asInstanceOf[Hole]
335363
}
336364

337365
override def transform(tree: Tree)(implicit ctx: Context): Tree =

tests/pos-macros/i7110a/Macro_1.scala

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,15 @@
1+
import scala.quoted._
2+
3+
object Macros {
4+
5+
inline def m[R](sym: Symantics[R]) : R = ${ mImpl[R]('{sym}) }
6+
7+
def mImpl[R: Type](using qctx: QuoteContext)(sym: Expr[Symantics[R]]): Expr[R] = '{
8+
$sym.Meth(42)
9+
}
10+
}
11+
12+
trait Symantics[R] {
13+
def Meth(exp: Int): R
14+
def Meth(): R
15+
}

tests/pos-macros/i7110a/Test_2.scala

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,14 @@
1+
import scala.quoted._
2+
import Macros._
3+
4+
object Test {
5+
def main(args: Array[String]): Unit = {
6+
7+
val sym = new Symantics[Int] {
8+
def Meth(exp: Int): Int = exp
9+
def Meth(): Int = 42
10+
}
11+
12+
val test = m[Int](sym)
13+
}
14+
}

tests/pos-macros/i7110b/Macro_1.scala

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,16 @@
1+
import scala.quoted._
2+
3+
object Macros {
4+
5+
inline def m[T](sym: Symantics {type R = T}) : T = ${ mImpl[T]('{sym}) }
6+
7+
def mImpl[T: Type](using qctx: QuoteContext)(sym: Expr[Symantics { type R = T }]): Expr[T] = '{
8+
$sym.Meth(42)
9+
}
10+
}
11+
12+
trait Symantics {
13+
type R
14+
def Meth(exp: Int): R
15+
def Meth(): R
16+
}

tests/pos-macros/i7110b/Test_2.scala

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,15 @@
1+
import scala.quoted._
2+
import Macros._
3+
4+
object Test {
5+
def main(args: Array[String]): Unit = {
6+
7+
val sym = new Symantics {
8+
type R = Int
9+
def Meth(exp: Int): Int = exp
10+
def Meth(): Int = 42
11+
}
12+
13+
val test = m(sym)
14+
}
15+
}

tests/pos-macros/i7110c/Macro_1.scala

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,14 @@
1+
import scala.quoted._
2+
3+
object Macros {
4+
5+
inline def m[R](sym: Symantics[R]) : R = ${ mImpl[R]('{sym}) }
6+
7+
def mImpl[R: Type](using qctx: QuoteContext)(sym: Expr[Symantics[R]]): Expr[R] = '{
8+
$sym.Meth(42)
9+
}
10+
}
11+
12+
trait Symantics[R] {
13+
def Meth(exp: Int): R
14+
}

tests/pos-macros/i7110c/Test_2.scala

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,16 @@
1+
import scala.quoted._
2+
import Macros._
3+
4+
object Test {
5+
def main(args: Array[String]): Unit = {
6+
7+
val sym = new Symantics2
8+
9+
val test = m[Int](sym)
10+
}
11+
}
12+
13+
class Symantics2 extends Symantics[Int] {
14+
def Meth(exp: Int): Int = exp
15+
def Meth(): Int = 42
16+
}

0 commit comments

Comments
 (0)