|
| 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