Skip to content
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.

Commit 50ae350

Browse files
committedJul 26, 2021
Fix Reflection wildcard abstraction
Currently, we are missing in the API the Wildcard concept for `case _ =>` patterns. These are encoded as term `Ident` with name `_`. This tree can be inconsistently matched by `Ident`, `TypeIdent` or `WildcardTypeTree`. There is also no way to create an `Ident(_)`. The typed expression can contain non-Term trees: Wildcard, Alternatives, Bind and Unapply. The solution is to add a TypedTree supertype of Typed that contains a Tree. Changes * `Ident` does not match `Ident(_)` * `TypeIdent` does not match `Ident(_)` * `WildcardTypeTree` does not match `Ident(_)` if it is a term * Add `Wildcard` type that matches a term `Ident(_)` * `Typed` only matched if the expr is a `Term` * Add `TypedTree` Fixes #12188
1 parent 7b4091b commit 50ae350

File tree

18 files changed

+256
-46
lines changed

18 files changed

+256
-46
lines changed
 

‎compiler/src/scala/quoted/runtime/impl/QuotesImpl.scala

Lines changed: 54 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -341,11 +341,19 @@ class QuotesImpl private (using val ctx: Context) extends Quotes, QuoteUnpickler
341341

342342
object TermTypeTest extends TypeTest[Tree, Term]:
343343
def unapply(x: Tree): Option[Term & x.type] = x match
344-
case x: tpd.PatternTree => None
344+
case x: (tpd.Ident & x.type) =>
345+
if x.isTerm && x.name != nme.WILDCARD then Some(x)
346+
else None
347+
case x: (tpd.Typed & x.type) =>
348+
// Matches `Typed` but not `TypedTree`
349+
TypedTypeTest.unapply(x)
345350
case x: (tpd.SeqLiteral & x.type) => Some(x)
346351
case x: (tpd.Inlined & x.type) => Some(x)
347352
case x: (tpd.NamedArg & x.type) => Some(x)
348-
case _ => if x.isTerm then Some(x) else None
353+
case x: tpd.PatternTree => None
354+
case _ =>
355+
if x.isTerm then Some(x)
356+
else None
349357
end TermTypeTest
350358

351359
object Term extends TermModule:
@@ -429,7 +437,7 @@ class QuotesImpl private (using val ctx: Context) extends Quotes, QuoteUnpickler
429437

430438
object IdentTypeTest extends TypeTest[Tree, Ident]:
431439
def unapply(x: Tree): Option[Ident & x.type] = x match
432-
case x: (tpd.Ident & x.type) if x.isTerm => Some(x)
440+
case x: (tpd.Ident & x.type) if x.isTerm && x.name != nme.WILDCARD => Some(x)
433441
case _ => None
434442
end IdentTypeTest
435443

@@ -655,7 +663,10 @@ class QuotesImpl private (using val ctx: Context) extends Quotes, QuoteUnpickler
655663

656664
object TypedTypeTest extends TypeTest[Tree, Typed]:
657665
def unapply(x: Tree): Option[Typed & x.type] = x match
658-
case x: (tpd.Typed & x.type) => Some(x)
666+
case x: (tpd.Typed & x.type) =>
667+
x.expr match
668+
case TermTypeTest(_) => Some(x)
669+
case _ => None
659670
case _ => None
660671
end TypedTypeTest
661672

@@ -1406,6 +1417,45 @@ class QuotesImpl private (using val ctx: Context) extends Quotes, QuoteUnpickler
14061417
end extension
14071418
end TypeCaseDefMethods
14081419

1420+
1421+
type Wildcard = tpd.Ident
1422+
1423+
object WildcardTypeTest extends TypeTest[Tree, Wildcard]:
1424+
def unapply(x: Tree): Option[Wildcard & x.type] = x match
1425+
case x: (tpd.Ident & x.type) if x.name == nme.WILDCARD => Some(x)
1426+
case _ => None
1427+
end WildcardTypeTest
1428+
1429+
object Wildcard extends WildcardModule:
1430+
def apply(): Wildcard =
1431+
withDefaultPos(untpd.Ident(nme.WILDCARD).withType(dotc.core.Symbols.defn.AnyType))
1432+
def unapply(pattern: Wildcard): true = true
1433+
end Wildcard
1434+
1435+
type TypedTree = tpd.Typed
1436+
1437+
object TypedTreeTypeTest extends TypeTest[Tree, TypedTree]:
1438+
def unapply(x: Tree): Option[TypedTree & x.type] = x match
1439+
case x: (tpd.Typed & x.type) => Some(x)
1440+
case _ => None
1441+
end TypedTreeTypeTest
1442+
1443+
object TypedTree extends TypedTreeModule:
1444+
def apply(expr: Term, tpt: TypeTree): Typed =
1445+
withDefaultPos(tpd.Typed(xCheckMacroValidExpr(expr), tpt))
1446+
def copy(original: Tree)(expr: Term, tpt: TypeTree): Typed =
1447+
tpd.cpy.Typed(original)(xCheckMacroValidExpr(expr), tpt)
1448+
def unapply(x: Typed): (Term, TypeTree) =
1449+
(x.expr, x.tpt)
1450+
end TypedTree
1451+
1452+
given TypedTreeMethods: TypedTreeMethods with
1453+
extension (self: Typed)
1454+
def tree: Tree = self.expr
1455+
def tpt: TypeTree = self.tpt
1456+
end extension
1457+
end TypedTreeMethods
1458+
14091459
type Bind = tpd.Bind
14101460

14111461
object BindTypeTest extends TypeTest[Tree, Bind]:

‎compiler/src/scala/quoted/runtime/impl/printers/Extractors.scala

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -166,12 +166,16 @@ object Extractors {
166166
this += "CaseDef(" += pat += ", " += guard += ", " += body += ")"
167167
case TypeCaseDef(pat, body) =>
168168
this += "TypeCaseDef(" += pat += ", " += body += ")"
169+
case Wildcard() =>
170+
this += "Wildcard()"
169171
case Bind(name, body) =>
170172
this += "Bind(\"" += name += "\", " += body += ")"
171173
case Unapply(fun, implicits, patterns) =>
172174
this += "Unapply(" += fun += ", " ++= implicits += ", " ++= patterns += ")"
173175
case Alternatives(patterns) =>
174176
this += "Alternatives(" ++= patterns += ")"
177+
case TypedTree(tree, tpt) =>
178+
this += "TypedTree(" += tree += ", " += tpt += ")"
175179
}
176180

177181
def visitConstant(x: Constant): this.type = x match {

‎compiler/src/scala/quoted/runtime/impl/printers/SourceCode.scala

Lines changed: 20 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -328,7 +328,7 @@ object SourceCode {
328328
}
329329
this
330330

331-
case Ident("_") =>
331+
case Wildcard() =>
332332
this += "_"
333333

334334
case tree: Ident =>
@@ -453,6 +453,15 @@ object SourceCode {
453453
printTypeOrAnnots(tpt.tpe)
454454
}
455455
}
456+
case TypedTree(tree1, tpt) =>
457+
printPattern(tree1)
458+
tree1 match
459+
case Wildcard() =>
460+
this += ":"
461+
printType(tpt.tpe)
462+
case _ => // Alternatives, Unapply, Bind
463+
this
464+
456465

457466
case Assign(lhs, rhs) =>
458467
printTree(lhs)
@@ -896,13 +905,13 @@ object SourceCode {
896905
}
897906

898907
private def printPattern(pattern: Tree): this.type = pattern match {
899-
case Ident("_") =>
908+
case Wildcard() =>
900909
this += "_"
901910

902-
case Bind(name, Ident("_")) =>
911+
case Bind(name, Wildcard()) =>
903912
this += name
904913

905-
case Bind(name, Typed(Ident("_"), tpt)) =>
914+
case Bind(name, Typed(Wildcard(), tpt)) =>
906915
this += highlightValDef(name) += ": "
907916
printTypeTree(tpt)
908917

@@ -928,9 +937,13 @@ object SourceCode {
928937
case Alternatives(trees) =>
929938
inParens(printPatterns(trees, " | "))
930939

931-
case Typed(Ident("_"), tpt) =>
932-
this += "_: "
933-
printTypeTree(tpt)
940+
case TypedTree(tree1, tpt) =>
941+
tree1 match
942+
case Wildcard() =>
943+
this += "_: "
944+
printTypeTree(tpt)
945+
case _ =>
946+
printPattern(tree1)
934947

935948
case v: Term =>
936949
printTree(v)

‎library/src/scala/quoted/Quotes.scala

Lines changed: 71 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -133,7 +133,6 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
133133
* | +- Apply
134134
* | +- TypeApply
135135
* | +- Super
136-
* | +- Typed
137136
* | +- Assign
138137
* | +- Block
139138
* | +- Closure
@@ -146,7 +145,16 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
146145
* | +- Inlined
147146
* | +- SelectOuter
148147
* | +- While
148+
* | +---+- Typed
149+
* | /
150+
* +- TypedTree +------------------·
151+
* +- Wildcard
152+
* +- Bind
153+
* +- Unapply
154+
* +- Alternatives
149155
* |
156+
* +- CaseDef
157+
* +- TypeCaseDef
150158
* |
151159
* +- TypeTree ----+- Inferred
152160
* | +- TypeIdent
@@ -164,13 +172,6 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
164172
* |
165173
* +- TypeBoundsTree
166174
* +- WildcardTypeTree
167-
* |
168-
* +- CaseDef
169-
* |
170-
* +- TypeCaseDef
171-
* +- Bind
172-
* +- Unapply
173-
* +- Alternatives
174175
*
175176
* +- ParamClause -+- TypeParamClause
176177
* +- TermParamClause
@@ -1120,8 +1121,12 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
11201121
/** `TypeTest` that allows testing at runtime in a pattern match if a `Tree` is a `Typed` */
11211122
given TypedTypeTest: TypeTest[Tree, Typed]
11221123

1123-
/** Tree representing a type ascription `x: T` in the source code */
1124-
type Typed <: Term
1124+
/** Tree representing a type ascription `x: T` in the source code.
1125+
*
1126+
* Also represents a pattern that contains a term `x`.
1127+
* Other `: T` patterns use the more general `TypeTree`.
1128+
*/
1129+
type Typed <: Term & TypeTree
11251130

11261131
/** Module object of `type Typed` */
11271132
val Typed: TypedModule
@@ -2049,6 +2054,56 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
20492054

20502055
// ----- Patterns ------------------------------------------------
20512056

2057+
/** Pattern representing a `_` wildcard. */
2058+
type Wildcard <: Tree
2059+
2060+
/** `TypeTest` that allows testing at runtime in a pattern match if a `Tree` is a `Wildcard` */
2061+
given WildcardTypeTest: TypeTest[Tree, Wildcard]
2062+
2063+
/** Module object of `type Wildcard` */
2064+
val Wildcard: WildcardModule
2065+
2066+
/** Methods of the module object `val Wildcard` */
2067+
trait WildcardModule { this: Wildcard.type =>
2068+
def apply(): Wildcard
2069+
def unapply(pattern: Wildcard): true
2070+
}
2071+
2072+
/** `TypeTest` that allows testing at runtime in a pattern match if a `Tree` is a `TypedTree` */
2073+
given TypedTreeTypeTest: TypeTest[Tree, TypedTree]
2074+
2075+
/** Tree representing a type ascription or pattern `x: T` in the source code
2076+
*
2077+
* The tree `x` may contain a `Constant`, `Ref`, `Wildcard`, `Bind`, `Unapply` or `Alternatives`.
2078+
*/
2079+
type TypedTree <: Term
2080+
2081+
/** Module object of `type TypedTree` */
2082+
val TypedTree: TypedTreeModule
2083+
2084+
/** Methods of the module object `val TypedTree` */
2085+
trait TypedTreeModule { this: TypedTree.type =>
2086+
2087+
/** Create a type ascription `<x: Tree>: <tpt: TypeTree>` */
2088+
def apply(expr: Tree, tpt: TypeTree): TypedTree
2089+
2090+
def copy(original: Tree)(expr: Tree, tpt: TypeTree): TypedTree
2091+
2092+
/** Matches `<expr: Tree>: <tpt: TypeTree>` */
2093+
def unapply(x: TypedTree): (Tree, TypeTree)
2094+
}
2095+
2096+
/** Makes extension methods on `TypedTree` available without any imports */
2097+
given TypedTreeMethods: TypedTreeMethods
2098+
2099+
/** Extension methods of `TypedTree` */
2100+
trait TypedTreeMethods:
2101+
extension (self: TypedTree)
2102+
def tree: Tree
2103+
def tpt: TypeTree
2104+
end extension
2105+
end TypedTreeMethods
2106+
20522107
/** Pattern representing a `_ @ _` binding. */
20532108
type Bind <: Tree
20542109

@@ -4327,9 +4382,11 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
43274382
case TypeBoundsTree(lo, hi) => foldTree(foldTree(x, lo)(owner), hi)(owner)
43284383
case CaseDef(pat, guard, body) => foldTree(foldTrees(foldTree(x, pat)(owner), guard)(owner), body)(owner)
43294384
case TypeCaseDef(pat, body) => foldTree(foldTree(x, pat)(owner), body)(owner)
4385+
case Wildcard() => x
43304386
case Bind(_, body) => foldTree(x, body)(owner)
43314387
case Unapply(fun, implicits, patterns) => foldTrees(foldTrees(foldTree(x, fun)(owner), implicits)(owner), patterns)(owner)
43324388
case Alternatives(patterns) => foldTrees(x, patterns)(owner)
4389+
case TypedTree(tree1, tpt) => foldTree(foldTree(x, tree1)(owner), tpt)(owner)
43334390
}
43344391
}
43354392
end TreeAccumulator
@@ -4387,12 +4444,15 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
43874444
transformCaseDef(tree)(owner)
43884445
case tree: TypeCaseDef =>
43894446
transformTypeCaseDef(tree)(owner)
4447+
case Wildcard() => tree
43904448
case pattern: Bind =>
43914449
Bind.copy(pattern)(pattern.name, pattern.pattern)
43924450
case pattern: Unapply =>
43934451
Unapply.copy(pattern)(transformTerm(pattern.fun)(owner), transformSubTrees(pattern.implicits)(owner), transformTrees(pattern.patterns)(owner))
43944452
case pattern: Alternatives =>
43954453
Alternatives.copy(pattern)(transformTrees(pattern.patterns)(owner))
4454+
case TypedTree(expr, tpt) =>
4455+
TypedTree.copy(tree)(transformTree(expr)(owner), transformTypeTree(tpt)(owner))
43964456
}
43974457
}
43984458

@@ -4443,7 +4503,7 @@ trait Quotes { self: runtime.QuoteUnpickler & runtime.QuoteMatching =>
44434503
case New(tpt) =>
44444504
New.copy(tree)(transformTypeTree(tpt)(owner))
44454505
case Typed(expr, tpt) =>
4446-
Typed.copy(tree)(/*FIXME #12222: transformTerm(expr)(owner)*/transformTree(expr)(owner).asInstanceOf[Term], transformTypeTree(tpt)(owner))
4506+
Typed.copy(tree)(transformTerm(expr)(owner), transformTypeTree(tpt)(owner))
44474507
case tree: NamedArg =>
44484508
NamedArg.copy(tree)(tree.name, transformTerm(tree.value)(owner))
44494509
case Assign(lhs, rhs) =>

‎project/MiMaFilters.scala

Lines changed: 14 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,19 @@ object MiMaFilters {
88
exclude[MissingClassProblem]("scala.annotation.internal.ProvisionalSuperClass"),
99

1010
// New APIs marked @experimental in 3.0.2
11-
exclude[MissingClassProblem]("scala.Selectable$WithoutPreciseParameterTypes")
11+
exclude[MissingClassProblem]("scala.Selectable$WithoutPreciseParameterTypes"),
12+
exclude[ReversedMissingMethodProblem]("scala.quoted.Quotes#reflectModule.WildcardTypeTest"),
13+
exclude[ReversedMissingMethodProblem]("scala.quoted.Quotes#reflectModule.Wildcard"),
14+
exclude[ReversedMissingMethodProblem]("scala.quoted.Quotes#reflectModule.TypedTreeTypeTest"),
15+
exclude[ReversedMissingMethodProblem]("scala.quoted.Quotes#reflectModule.TypedTree"),
16+
exclude[ReversedMissingMethodProblem]("scala.quoted.Quotes#reflectModule.TypedTreeMethods"),
17+
exclude[DirectMissingMethodProblem]("scala.quoted.Quotes#reflectModule.WildcardTypeTest"),
18+
exclude[DirectMissingMethodProblem]("scala.quoted.Quotes#reflectModule.Wildcard"),
19+
exclude[DirectMissingMethodProblem]("scala.quoted.Quotes#reflectModule.TypedTreeTypeTest"),
20+
exclude[DirectMissingMethodProblem]("scala.quoted.Quotes#reflectModule.TypedTree"),
21+
exclude[DirectMissingMethodProblem]("scala.quoted.Quotes#reflectModule.TypedTreeMethods"),
22+
exclude[MissingClassProblem]("scala.quoted.Quotes$reflectModule$TypedTreeMethods"),
23+
exclude[MissingClassProblem]("scala.quoted.Quotes$reflectModule$TypedTreeModule"),
24+
exclude[MissingClassProblem]("scala.quoted.Quotes$reflectModule$WildcardModule"),
1225
)
1326
}

‎tests/pos-macros/i11401/X_1.scala

Lines changed: 9 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -16,7 +16,7 @@ class SLSelect[S]:
1616

1717
def fold[S](s0:S)(step: (S,SLSelect[S])=> S): S = {
1818
???
19-
}
19+
}
2020

2121
def fold_async[S](s0:S)(step: (S,SLSelect[S])=> Future[S]): Future[S] = {
2222
???
@@ -27,7 +27,7 @@ class SLSelect[S]:
2727
await(s0.onRead(ch)(f).runAsync())
2828

2929
def runAsync(): Future[S] = ???
30-
30+
3131

3232

3333
object X:
@@ -36,30 +36,30 @@ object X:
3636
processImpl[T]('f)
3737
}
3838

39-
def processImpl[T:Type](t:Expr[T])(using Quotes):Expr[Future[T]] =
39+
def processImpl[T:Type](t:Expr[T])(using Quotes):Expr[Future[T]] =
4040
import quotes.reflect._
4141
val r = processTree[T](t.asTerm)
4242
r.asExprOf[Future[T]]
4343

44-
45-
def processTree[T:Type](using Quotes)(t: quotes.reflect.Term):quotes.reflect.Term =
44+
45+
def processTree[T:Type](using Quotes)(t: quotes.reflect.Term):quotes.reflect.Term =
4646
import quotes.reflect._
4747
val r: Term = t match
4848
case Inlined(_,List(),body) => processTree(body)
49-
case Inlined(d,bindings,body) =>
49+
case Inlined(d,bindings,body) =>
5050
Inlined(d,bindings,processTree[T](body))
5151
case Block(stats,expr) => Block(stats,processTree(expr))
5252
case Apply(Apply(TypeApply(Select(x,"fold"),targs),List(state)),List(fun)) =>
53-
val nFun = processLambda[T](fun)
53+
val nFun = processLambda[T](fun)
5454
Apply(Apply(TypeApply(Select.unique(x,"fold_async"),targs),List(state)),List(nFun))
5555
case Apply(TypeApply(Ident("await"),targs),List(body)) => body
5656
case Typed(x,tp) => Typed(processTree(x), Inferred(TypeRepr.of[Future].appliedTo(tp.tpe)) )
5757
case _ => throw new RuntimeException(s"tree not recoginized: $t")
5858
val checker = new TreeMap() {}
5959
checker.transformTerm(r)(Symbol.spliceOwner)
6060
r
61-
62-
def processLambda[T:Type](using Quotes)(fun: quotes.reflect.Term):quotes.reflect.Term =
61+
62+
def processLambda[T:Type](using Quotes)(fun: quotes.reflect.Term):quotes.reflect.Term =
6363
import quotes.reflect._
6464

6565
def changeArgs(oldArgs:List[Tree], newArgs:List[Tree], body:Term, owner: Symbol):Term =
Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,13 @@
1+
import scala.quoted.*
2+
3+
object MatchTest {
4+
inline def test[T](inline obj: Any): Unit = ${testImpl('obj)}
5+
6+
def testImpl[T](objExpr: Expr[T])(using Quotes): Expr[Unit] = {
7+
import quotes.reflect.*
8+
// test that the extractors work
9+
val Inlined(None, Nil, Block(Nil, Match(param @ Ident("a"), List(CaseDef(Literal(IntConstant(1)), None, Block(Nil, Literal(UnitConstant()))), CaseDef(Wildcard(), None, Block(Nil, Literal(UnitConstant()))))))) = objExpr.asTerm
10+
// test that the constructors work
11+
Block(Nil, Match(param, List(CaseDef(Literal(IntConstant(1)), None, Block(Nil, Literal(UnitConstant()))), CaseDef(Wildcard(), None, Block(Nil, Literal(UnitConstant())))))).asExprOf[Unit]
12+
}
13+
}

‎tests/pos-macros/i12188b/Test_2.scala

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,6 @@
1+
2+
def test(a: Int) = MatchTest.test {
3+
a match
4+
case 1 =>
5+
case _ =>
6+
}
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 MatchTest {
4+
inline def test(a: Int): Unit = ${testImpl('a)}
5+
6+
def testImpl(a: Expr[Any])(using Quotes): Expr[Unit] = {
7+
import quotes.reflect.*
8+
val matchTree = Match(a.asTerm, List(
9+
CaseDef(Literal(IntConstant(1)), None, Block(Nil, Literal(UnitConstant()))),
10+
CaseDef(Alternatives(List(Literal(IntConstant(2)), Literal(IntConstant(3)), Literal(IntConstant(4)))), None, Block(Nil, Literal(UnitConstant()))),
11+
CaseDef(TypedTree(Alternatives(List(Literal(IntConstant(4)), Literal(IntConstant(5)))), TypeIdent(defn.IntClass)), None, Block(Nil, Literal(UnitConstant()))),
12+
CaseDef(TypedTree(Wildcard(), TypeIdent(defn.IntClass)), None, Block(Nil, Literal(UnitConstant())))))
13+
matchTree.asExprOf[Unit]
14+
}
15+
}

‎tests/pos-macros/i12188c/Test_2.scala

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,2 @@
1+
2+
def test(a: Int) = MatchTest.test(a)
Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,3 +1,3 @@
11
DefDef("foo", Nil, TypeIdent("Int"), Some(Apply(Select(Literal(IntConstant(1)), "+"), List(Literal(IntConstant(2))))))
22
ValDef("bar", TypeIdent("Int"), Some(Apply(Select(Literal(IntConstant(2)), "+"), List(Literal(IntConstant(3))))))
3-
Bind("x", Ident("_"))
3+
Bind("x", Wildcard())
Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,3 +1,3 @@
11
DefDef("foo", Nil, TypeIdent("Int"), Some(Apply(Select(Literal(IntConstant(1)), "+"), List(Literal(IntConstant(2))))))
22
ValDef("bar", TypeIdent("Int"), Some(Apply(Select(Literal(IntConstant(2)), "+"), List(Literal(IntConstant(3))))))
3-
Bind("x", Ident("_"))
3+
Bind("x", Wildcard())

‎tests/run-macros/i12188.check

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,3 @@
1+
PC1
2+
PC2
3+
default

‎tests/run-macros/i12188/Macro_1.scala

Lines changed: 23 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,23 @@
1+
import scala.quoted.*
2+
3+
object MatchTest {
4+
inline def test[T](inline obj: T): String = ${testImpl('obj)}
5+
6+
def testImpl[T](objExpr: Expr[T])(using qctx: Quotes, t: Type[T]): Expr[String] = {
7+
import qctx.reflect.*
8+
9+
val obj = objExpr.asTerm
10+
val cases = obj.tpe.typeSymbol.children.map { child =>
11+
val subtype = TypeIdent(child)
12+
val bind = Symbol.newBind(Symbol.spliceOwner, "c", Flags.EmptyFlags, subtype.tpe)
13+
CaseDef(Bind(bind, Typed(Ref(bind), subtype)), None, Literal(StringConstant(subtype.show)))
14+
} ::: {
15+
CaseDef(Wildcard(), None, Literal(StringConstant("default")))
16+
} :: Nil
17+
val bind = Symbol.newBind(Symbol.spliceOwner, "o", Flags.EmptyFlags, obj.tpe)
18+
val result = Match(obj, cases)
19+
val code = result.show(using Printer.TreeAnsiCode)
20+
// println(code)
21+
result.asExprOf[String]
22+
}
23+
}

‎tests/run-macros/i12188/Test_2.scala

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,8 @@
1+
sealed trait P
2+
case class PC1(a: String) extends P
3+
case class PC2(b: Int) extends P
4+
5+
@main def Test =
6+
println(MatchTest.test(PC1("ab"): P))
7+
println(MatchTest.test(PC2(10): P))
8+
println(MatchTest.test(null: P))

‎tests/run-macros/tasty-extractors-1.check

Lines changed: 10 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -40,37 +40,37 @@ OrType(ConstantType(IntConstant(1)), ConstantType(IntConstant(2)))
4040
Inlined(None, Nil, Match(Literal(StringConstant("a")), List(CaseDef(Literal(StringConstant("a")), None, Block(Nil, Literal(UnitConstant()))))))
4141
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
4242

43-
Inlined(None, Nil, Match(Literal(StringConstant("b")), List(CaseDef(Bind("n", Ident("_")), None, Block(Nil, Literal(UnitConstant()))))))
43+
Inlined(None, Nil, Match(Literal(StringConstant("b")), List(CaseDef(Bind("n", Wildcard()), None, Block(Nil, Literal(UnitConstant()))))))
4444
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
4545

46-
Inlined(None, Nil, Match(Literal(StringConstant("c")), List(CaseDef(Bind("n", Typed(Ident("_"), TypeIdent("String"))), None, Block(Nil, Literal(UnitConstant()))))))
46+
Inlined(None, Nil, Match(Literal(StringConstant("c")), List(CaseDef(Bind("n", TypedTree(Wildcard(), TypeIdent("String"))), None, Block(Nil, Literal(UnitConstant()))))))
4747
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
4848

49-
Inlined(None, Nil, Match(Literal(StringConstant("e")), List(CaseDef(Ident("_"), None, Block(Nil, Literal(UnitConstant()))))))
49+
Inlined(None, Nil, Match(Literal(StringConstant("e")), List(CaseDef(Wildcard(), None, Block(Nil, Literal(UnitConstant()))))))
5050
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
5151

52-
Inlined(None, Nil, Match(Literal(StringConstant("f")), List(CaseDef(Typed(Ident("_"), TypeIdent("String")), None, Block(Nil, Literal(UnitConstant()))))))
52+
Inlined(None, Nil, Match(Literal(StringConstant("f")), List(CaseDef(TypedTree(Wildcard(), TypeIdent("String")), None, Block(Nil, Literal(UnitConstant()))))))
5353
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
5454

55-
Inlined(None, Nil, Match(Typed(Literal(StringConstant("g")), TypeIdent("Any")), List(CaseDef(Alternatives(List(Typed(Ident("_"), TypeIdent("String")), Typed(Ident("_"), TypeIdent("Int")))), None, Block(Nil, Literal(UnitConstant()))))))
55+
Inlined(None, Nil, Match(Typed(Literal(StringConstant("g")), TypeIdent("Any")), List(CaseDef(Alternatives(List(TypedTree(Wildcard(), TypeIdent("String")), TypedTree(Wildcard(), TypeIdent("Int")))), None, Block(Nil, Literal(UnitConstant()))))))
5656
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
5757

58-
Inlined(None, Nil, Match(Literal(StringConstant("h")), List(CaseDef(Ident("_"), Some(Literal(BooleanConstant(false))), Block(Nil, Literal(UnitConstant()))))))
58+
Inlined(None, Nil, Match(Literal(StringConstant("h")), List(CaseDef(Wildcard(), Some(Literal(BooleanConstant(false))), Block(Nil, Literal(UnitConstant()))))))
5959
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
6060

61-
Inlined(None, Nil, Block(List(ValDef("a", Inferred(), Some(Literal(StringConstant("o"))))), Match(Literal(StringConstant("i")), List(CaseDef(Bind("a", Ident("_")), None, Block(Nil, Literal(UnitConstant())))))))
61+
Inlined(None, Nil, Block(List(ValDef("a", Inferred(), Some(Literal(StringConstant("o"))))), Match(Literal(StringConstant("i")), List(CaseDef(Bind("a", Wildcard()), None, Block(Nil, Literal(UnitConstant())))))))
6262
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
6363

64-
Inlined(None, Nil, Match(Ident("Nil"), List(CaseDef(Unapply(TypeApply(Select(Ident("List"), "unapplySeq"), List(Inferred())), Nil, List(Bind("a", Ident("_")), Bind("b", Ident("_")), Bind("c", Ident("_")))), None, Block(Nil, Literal(UnitConstant()))))))
64+
Inlined(None, Nil, Match(Ident("Nil"), List(CaseDef(Unapply(TypeApply(Select(Ident("List"), "unapplySeq"), List(Inferred())), Nil, List(Bind("a", Wildcard()), Bind("b", Wildcard()), Bind("c", Wildcard()))), None, Block(Nil, Literal(UnitConstant()))))))
6565
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
6666

67-
Inlined(None, Nil, Try(Literal(IntConstant(1)), List(CaseDef(Ident("_"), None, Block(Nil, Literal(UnitConstant())))), None))
67+
Inlined(None, Nil, Try(Literal(IntConstant(1)), List(CaseDef(Wildcard(), None, Block(Nil, Literal(UnitConstant())))), None))
6868
OrType(ConstantType(IntConstant(1)), TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit"))
6969

7070
Inlined(None, Nil, Try(Literal(IntConstant(2)), Nil, Some(Literal(UnitConstant()))))
7171
ConstantType(IntConstant(2))
7272

73-
Inlined(None, Nil, Try(Literal(IntConstant(3)), List(CaseDef(Ident("_"), None, Block(Nil, Literal(UnitConstant())))), Some(Literal(UnitConstant()))))
73+
Inlined(None, Nil, Try(Literal(IntConstant(3)), List(CaseDef(Wildcard(), None, Block(Nil, Literal(UnitConstant())))), Some(Literal(UnitConstant()))))
7474
OrType(ConstantType(IntConstant(3)), TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit"))
7575

7676
Inlined(None, Nil, Literal(BooleanConstant(false)))

‎tests/run-macros/tasty-extractors-2.check

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -49,7 +49,7 @@ TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
4949
Inlined(None, Nil, Block(List(ClassDef("Foo", DefDef("<init>", List(TermParamClause(Nil)), Inferred(), None), List(Apply(Select(New(Inferred()), "<init>"), Nil)), None, List(DefDef("a", Nil, Inferred(), Some(Literal(IntConstant(0))))))), Literal(UnitConstant())))
5050
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
5151

52-
Inlined(None, Nil, Block(List(ClassDef("Foo", DefDef("<init>", List(TermParamClause(Nil)), Inferred(), None), List(Apply(Select(New(Inferred()), "<init>"), Nil), TypeSelect(Select(Ident("_root_"), "scala"), "Product"), TypeSelect(Select(Ident("_root_"), "scala"), "Serializable")), None, List(DefDef("hashCode", List(TermParamClause(Nil)), Inferred(), Some(Apply(Ident("_hashCode"), List(This(Some("Foo")))))), DefDef("equals", List(TermParamClause(List(ValDef("x$0", Inferred(), None)))), Inferred(), Some(Apply(Select(Apply(Select(This(Some("Foo")), "eq"), List(TypeApply(Select(Ident("x$0"), "$asInstanceOf$"), List(Inferred())))), "||"), List(Match(Ident("x$0"), List(CaseDef(Bind("x$0", Typed(Ident("_"), Inferred())), None, Apply(Select(Literal(BooleanConstant(true)), "&&"), List(Apply(Select(Ident("x$0"), "canEqual"), List(This(Some("Foo"))))))), CaseDef(Ident("_"), None, Literal(BooleanConstant(false))))))))), DefDef("toString", List(TermParamClause(Nil)), Inferred(), Some(Apply(Ident("_toString"), List(This(Some("Foo")))))), DefDef("canEqual", List(TermParamClause(List(ValDef("that", Inferred(), None)))), Inferred(), Some(TypeApply(Select(Ident("that"), "isInstanceOf"), List(Inferred())))), DefDef("productArity", Nil, Inferred(), Some(Literal(IntConstant(0)))), DefDef("productPrefix", Nil, Inferred(), Some(Literal(StringConstant("Foo")))), DefDef("productElement", List(TermParamClause(List(ValDef("n", Inferred(), None)))), Inferred(), Some(Match(Ident("n"), List(CaseDef(Ident("_"), None, Apply(Ident("throw"), List(Apply(Select(New(Inferred()), "<init>"), List(Apply(Select(Ident("n"), "toString"), Nil)))))))))), DefDef("productElementName", List(TermParamClause(List(ValDef("n", Inferred(), None)))), Inferred(), Some(Match(Ident("n"), List(CaseDef(Ident("_"), None, Apply(Ident("throw"), List(Apply(Select(New(Inferred()), "<init>"), List(Apply(Select(Ident("n"), "toString"), Nil)))))))))), DefDef("copy", List(TermParamClause(Nil)), Inferred(), Some(Apply(Select(New(Inferred()), "<init>"), Nil))))), ValDef("Foo", TypeIdent("Foo$"), Some(Apply(Select(New(TypeIdent("Foo$")), "<init>"), Nil))), ClassDef("Foo$", DefDef("<init>", List(TermParamClause(Nil)), Inferred(), None), List(Apply(Select(New(Inferred()), "<init>"), Nil), Inferred()), Some(ValDef("_", Singleton(Ident("Foo")), None)), List(DefDef("apply", List(TermParamClause(Nil)), Inferred(), Some(Apply(Select(New(Inferred()), "<init>"), Nil))), DefDef("unapply", List(TermParamClause(List(ValDef("x$1", Inferred(), None)))), Singleton(Literal(BooleanConstant(true))), Some(Literal(BooleanConstant(true)))), DefDef("toString", Nil, Inferred(), Some(Literal(StringConstant("Foo")))), TypeDef("MirroredMonoType", TypeBoundsTree(Inferred(), Inferred())), DefDef("fromProduct", List(TermParamClause(List(ValDef("x$0", Inferred(), None)))), Inferred(), Some(Apply(Select(New(Inferred()), "<init>"), Nil)))))), Literal(UnitConstant())))
52+
Inlined(None, Nil, Block(List(ClassDef("Foo", DefDef("<init>", List(TermParamClause(Nil)), Inferred(), None), List(Apply(Select(New(Inferred()), "<init>"), Nil), TypeSelect(Select(Ident("_root_"), "scala"), "Product"), TypeSelect(Select(Ident("_root_"), "scala"), "Serializable")), None, List(DefDef("hashCode", List(TermParamClause(Nil)), Inferred(), Some(Apply(Ident("_hashCode"), List(This(Some("Foo")))))), DefDef("equals", List(TermParamClause(List(ValDef("x$0", Inferred(), None)))), Inferred(), Some(Apply(Select(Apply(Select(This(Some("Foo")), "eq"), List(TypeApply(Select(Ident("x$0"), "$asInstanceOf$"), List(Inferred())))), "||"), List(Match(Ident("x$0"), List(CaseDef(Bind("x$0", TypedTree(Wildcard(), Inferred())), None, Apply(Select(Literal(BooleanConstant(true)), "&&"), List(Apply(Select(Ident("x$0"), "canEqual"), List(This(Some("Foo"))))))), CaseDef(Wildcard(), None, Literal(BooleanConstant(false))))))))), DefDef("toString", List(TermParamClause(Nil)), Inferred(), Some(Apply(Ident("_toString"), List(This(Some("Foo")))))), DefDef("canEqual", List(TermParamClause(List(ValDef("that", Inferred(), None)))), Inferred(), Some(TypeApply(Select(Ident("that"), "isInstanceOf"), List(Inferred())))), DefDef("productArity", Nil, Inferred(), Some(Literal(IntConstant(0)))), DefDef("productPrefix", Nil, Inferred(), Some(Literal(StringConstant("Foo")))), DefDef("productElement", List(TermParamClause(List(ValDef("n", Inferred(), None)))), Inferred(), Some(Match(Ident("n"), List(CaseDef(Wildcard(), None, Apply(Ident("throw"), List(Apply(Select(New(Inferred()), "<init>"), List(Apply(Select(Ident("n"), "toString"), Nil)))))))))), DefDef("productElementName", List(TermParamClause(List(ValDef("n", Inferred(), None)))), Inferred(), Some(Match(Ident("n"), List(CaseDef(Wildcard(), None, Apply(Ident("throw"), List(Apply(Select(New(Inferred()), "<init>"), List(Apply(Select(Ident("n"), "toString"), Nil)))))))))), DefDef("copy", List(TermParamClause(Nil)), Inferred(), Some(Apply(Select(New(Inferred()), "<init>"), Nil))))), ValDef("Foo", TypeIdent("Foo$"), Some(Apply(Select(New(TypeIdent("Foo$")), "<init>"), Nil))), ClassDef("Foo$", DefDef("<init>", List(TermParamClause(Nil)), Inferred(), None), List(Apply(Select(New(Inferred()), "<init>"), Nil), Inferred()), Some(ValDef("_", Singleton(Ident("Foo")), None)), List(DefDef("apply", List(TermParamClause(Nil)), Inferred(), Some(Apply(Select(New(Inferred()), "<init>"), Nil))), DefDef("unapply", List(TermParamClause(List(ValDef("x$1", Inferred(), None)))), Singleton(Literal(BooleanConstant(true))), Some(Literal(BooleanConstant(true)))), DefDef("toString", Nil, Inferred(), Some(Literal(StringConstant("Foo")))), TypeDef("MirroredMonoType", TypeBoundsTree(Inferred(), Inferred())), DefDef("fromProduct", List(TermParamClause(List(ValDef("x$0", Inferred(), None)))), Inferred(), Some(Apply(Select(New(Inferred()), "<init>"), Nil)))))), Literal(UnitConstant())))
5353
TypeRef(ThisType(TypeRef(NoPrefix(), "scala")), "Unit")
5454

5555
Inlined(None, Nil, Block(List(ClassDef("Foo1", DefDef("<init>", List(TermParamClause(List(ValDef("a", TypeIdent("Int"), None)))), Inferred(), None), List(Apply(Select(New(Inferred()), "<init>"), Nil)), None, List(ValDef("a", Inferred(), None)))), Literal(UnitConstant())))

‎tests/run-staging/i5161.check

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,6 +1,6 @@
11
run : Some(2)
22
show : scala.Tuple2.apply[scala.Option[scala.Int], scala.Option[scala.Int]](scala.Some.apply[scala.Int](1), scala.Some.apply[scala.Int](1)) match {
3-
case scala.Tuple2((scala.Some(x): scala.Some[scala.Int]), (scala.Some(y): scala.Some[scala.Int])) =>
3+
case scala.Tuple2(scala.Some(x), scala.Some(y)) =>
44
scala.Some.apply[scala.Int](x.+(y))
55
case _ =>
66
scala.None

0 commit comments

Comments
 (0)
Please sign in to comment.