Skip to content

Commit

Permalink
feat: allow interpolation of fvars, with a warning
Browse files Browse the repository at this point in the history
  • Loading branch information
eric-wieser committed Nov 2, 2024
1 parent b7ff016 commit f95d85e
Show file tree
Hide file tree
Showing 2 changed files with 35 additions and 0 deletions.
4 changes: 4 additions & 0 deletions Qq/Macro.lean
Original file line number Diff line number Diff line change
Expand Up @@ -253,6 +253,10 @@ partial def unquoteExpr (e : Expr) : UnquoteM Expr := do
abstractedFVars := s.abstractedFVars.push fvarId
}
return fv
if let .fvar _ := e then
Lean.logWarning m!"Interpolated variable `{e}` should have type `Q(_)` not `{eTy}`. \
Adding `have {e} : Q(_) := {e}` may fix this."
return e
let e ← whnf e
let .const c _ := e.getAppFn | throwError "unquoteExpr: {e} : {eTy}"
let nargs := e.getAppNumArgs
Expand Down
31 changes: 31 additions & 0 deletions examples/unquoteExprWarning.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,31 @@
import Lean
import Qq.Macro

open Qq
open Lean Meta Elab Tactic Command

elab "iff_constructor" : tactic => do
let mvarId ← getMainGoal
let goalType ← getMainTarget

let_expr Iff a b := goalType
| throwError "Goal type is not of the form `a ↔ b`"

let a : Q(Prop) := a
let b : Q(Prop) := b

let mvarId1 ← mkFreshExprMVar q($a → $b) (userName := `mp)
let mvarId2 ← mkFreshExprMVar q($b → $a) (userName := `mpr)

let mvarId1 : Q($a → $b) := mvarId1
let mvarId2 : Q($b → $a) := mvarId2

mvarId.assign q(@Iff.intro $a $b $mvarId1 $mvarId2)

modify fun _ => { goals := [mvarId1.mvarId!, mvarId2.mvarId!] }

example (p q : Prop) (h1 : p → q) (h2 : q → p) : p ↔ q := by
iff_constructor

case mp => assumption
case mpr => assumption

0 comments on commit f95d85e

Please sign in to comment.