Skip to content

Commit c59b95e

Browse files
eric-wieseruniwuni
authored andcommitted
feat: add an eval% elaborator for interpolating the output of #eval (#10742)
This is taken from [my zulip message here](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/How.20to.20simplify.20this.20proof.20without.20using.20a.20have.20statement.3F/near/422294189). As an example: ```lean example : 2^10 = eval% 2^10 := by -- goal is `2^10 = 1024` sorry ```
1 parent e1337ab commit c59b95e

File tree

5 files changed

+99
-0
lines changed

5 files changed

+99
-0
lines changed

Mathlib.lean

+1
Original file line numberDiff line numberDiff line change
@@ -3526,6 +3526,7 @@ import Mathlib.Tactic.DeriveFintype
35263526
import Mathlib.Tactic.DeriveToExpr
35273527
import Mathlib.Tactic.DeriveTraversable
35283528
import Mathlib.Tactic.Eqns
3529+
import Mathlib.Tactic.Eval
35293530
import Mathlib.Tactic.Existsi
35303531
import Mathlib.Tactic.Explode
35313532
import Mathlib.Tactic.Explode.Datatypes

Mathlib/Tactic.lean

+1
Original file line numberDiff line numberDiff line change
@@ -43,6 +43,7 @@ import Mathlib.Tactic.DeriveFintype
4343
import Mathlib.Tactic.DeriveToExpr
4444
import Mathlib.Tactic.DeriveTraversable
4545
import Mathlib.Tactic.Eqns
46+
import Mathlib.Tactic.Eval
4647
import Mathlib.Tactic.Existsi
4748
import Mathlib.Tactic.Explode
4849
import Mathlib.Tactic.Explode.Datatypes

Mathlib/Tactic/Eval.lean

+45
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,45 @@
1+
/-
2+
Copyright (c) 2024 Eric Wieser. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Eric Wieser
5+
-/
6+
7+
import Qq.Macro
8+
9+
/-!
10+
# The `eval%` term elaborator
11+
12+
This file provides the `eval% x` term elaborator, which evaluates the constant `x` at compile-time
13+
in the interpreter, and interpolates it into the expression.
14+
-/
15+
16+
17+
open Qq Lean Elab Term
18+
19+
namespace Mathlib.Meta
20+
21+
/--
22+
`eval% x` evaluates the term `x : X` in the interpreter, and then injects the resulting expression.
23+
24+
As an example:
25+
```lean
26+
example : 2^10 = eval% 2^10 := by
27+
-- goal is `2^10 = 1024`
28+
sorry
29+
```
30+
This only works if a `Lean.ToExpr X` instance is available.
31+
32+
Tip: you can use `show_term eval% x` to see the value of `eval% x`.
33+
-/
34+
syntax (name := eval_expr) "eval% " term : term
35+
36+
@[term_elab eval_expr, inherit_doc eval_expr]
37+
unsafe def elabEvalExpr : Lean.Elab.Term.TermElab
38+
| `(term| eval% $stx) => fun exp => do
39+
let e ← Lean.Elab.Term.elabTermAndSynthesize stx exp
40+
let e ← instantiateMVars e
41+
let ee ← Meta.mkAppM ``Lean.toExpr #[e]
42+
Lean.Meta.evalExpr Expr q(Expr) ee (safety := .unsafe)
43+
| _ => fun _ => Elab.throwUnsupportedSyntax
44+
45+
end Mathlib.Meta

scripts/noshake.json

+1
Original file line numberDiff line numberDiff line change
@@ -268,6 +268,7 @@
268268
"Mathlib.Tactic.FunProp",
269269
"Mathlib.Tactic.FunProp.Measurable"],
270270
"Mathlib.Tactic.FinCases": ["Mathlib.Data.Fintype.Basic"],
271+
"Mathlib.Tactic.Eval": ["Qq.Macro"],
271272
"Mathlib.Tactic.DeriveFintype":
272273
["Mathlib.Data.Fintype.Basic",
273274
"Mathlib.Data.Fintype.Sigma",

test/eval_elab.lean

+51
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,51 @@
1+
import Mathlib.Tactic.Eval
2+
import Mathlib.Data.Finset.Powerset
3+
import Mathlib.Data.Finset.Sort
4+
5+
#guard_expr eval% 2^10 =ₛ 1024
6+
7+
#guard_expr (eval% 2^10 : Int) =ₛ .ofNat 1024
8+
9+
-- https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/How.20to.20simplify.20this.20proof.20without.20using.20a.20have.20statement.3F/near/422294189
10+
section from_zulip
11+
12+
/--
13+
error: failed to synthesize
14+
Lean.ToExpr (Finset (Finset ℕ))
15+
-/
16+
#guard_msgs in
17+
#check eval% Finset.powerset ({1, 2, 3} : Finset ℕ)
18+
19+
open Lean Qq
20+
21+
/-- `HasInstance (Foo X)` means than an `inst : Foo X` exists. -/
22+
class HasInstance (α : Type u) where
23+
/-- The reflected version of `inst`. -/
24+
expr : Expr
25+
26+
-- this obviously doesn't scale, which is why this is only in the test file
27+
instance : HasInstance (DecidableEq ℕ) :=
28+
⟨q(inferInstanceAs <| DecidableEq ℕ)⟩
29+
instance : HasInstance (DecidableEq (Finset ℕ)) :=
30+
⟨q(inferInstanceAs <| DecidableEq (Finset ℕ))⟩
31+
instance : HasInstance (DecidableEq (Finset (Finset ℕ))) :=
32+
⟨q(inferInstanceAs <| DecidableEq (Finset (Finset ℕ)))⟩
33+
34+
open Qq Lean
35+
/-- `Finset α` can be converted to an expr only if there is some way to find `DecidableEq α`. -/
36+
unsafe nonrec instance Finset.toExpr
37+
{α : Type u} [ToLevel.{u}] [ToExpr α] [HasInstance (DecidableEq α)] : ToExpr (Finset α) :=
38+
have u' : Level := ToLevel.toLevel.{u}
39+
have α' : Q(Type u') := Lean.ToExpr.toTypeExpr α
40+
letI : Q(DecidableEq $α') := HasInstance.expr (DecidableEq α)
41+
{ toTypeExpr := q(Finset $α')
42+
toExpr := fun x => show Q(Finset $α') from
43+
match show List Q($α') from x.val.unquot.reverse.map toExpr with
44+
| [] => q(∅)
45+
| x0 :: xs => List.foldl (fun s x => q(insert $x $s)) q({$x0}) xs }
46+
47+
#guard_expr
48+
(eval% Finset.powerset ({1, 2, 3} : Finset ℕ)) =
49+
{∅, {1}, {2}, {1, 2}, {3}, {1, 3}, {2, 3}, {1, 2, 3}}
50+
51+
end from_zulip

0 commit comments

Comments
 (0)