Skip to content

Commit

Permalink
feat: make intro be aware of let_fun (#3115)
Browse files Browse the repository at this point in the history
Adds support for `let_fun` to the `intro` and `intros` tactics. Also
adds support to `intro` for anonymous binder names, since the default
variable name for a `letFun` with an eta reduced body is anonymous.
  • Loading branch information
kmill authored Jan 31, 2024
1 parent dd77dbd commit 3198109
Show file tree
Hide file tree
Showing 2 changed files with 62 additions and 18 deletions.
54 changes: 36 additions & 18 deletions src/Lean/Meta/Tactic/Intro.lean
Original file line number Diff line number Diff line change
Expand Up @@ -44,24 +44,35 @@ namespace Lean.Meta
let fvars := fvars.push fvar
loop i lctx fvars j s body
| i+1, type =>
let type := type.instantiateRevRange j fvars.size fvars
withReader (fun ctx => { ctx with lctx := lctx }) do
withNewLocalInstances fvars j do
/- We used to use just `whnf`, but it produces counterintuitive behavior if
- `type` is a metavariable `?m` such that `?m := let x := v; b`, or
- `type` has `MData` or annotations such as `optParam` around a `let`-expression.
`whnf` instantiates metavariables, and consumes `MData`, but it also expands the `let`.
-/
let newType := (← instantiateMVars type).cleanupAnnotations
if newType.isForall || newType.isLet then
loop (i+1) lctx fvars fvars.size s newType
else
let newType ← whnf newType
if newType.isForall then
if let some (n, type, val, body) := type.letFun? then
let type := type.instantiateRevRange j fvars.size fvars
let type := type.headBeta
let val := val.instantiateRevRange j fvars.size fvars
let fvarId ← mkFreshFVarId
let (n, s) ← mkName lctx n true s
let lctx := lctx.mkLetDecl fvarId n type val
let fvar := mkFVar fvarId
let fvars := fvars.push fvar
loop i lctx fvars j s body
else
let type := type.instantiateRevRange j fvars.size fvars
withReader (fun ctx => { ctx with lctx := lctx }) do
withNewLocalInstances fvars j do
/- We used to use just `whnf`, but it produces counterintuitive behavior if
- `type` is a metavariable `?m` such that `?m := let x := v; b`, or
- `type` has `MData` or annotations such as `optParam` around a `let`-expression.
`whnf` instantiates metavariables, and consumes `MData`, but it also expands the `let`.
-/
let newType := (← instantiateMVars type).cleanupAnnotations
if newType.isForall || newType.isLet || newType.isLetFun then
loop (i+1) lctx fvars fvars.size s newType
else
throwTacticEx `introN mvarId "insufficient number of binders"
let newType ← whnf newType
if newType.isForall then
loop (i+1) lctx fvars fvars.size s newType
else
throwTacticEx `introN mvarId "insufficient number of binders"
let (fvars, mvarId) ← loop n lctx #[] 0 s mvarType
return (fvars.map Expr.fvarId!, mvarId)

Expand Down Expand Up @@ -97,6 +108,9 @@ private def mkAuxNameImp (preserveBinderNames : Bool) (hygienic : Bool) (useName
mkAuxNameWithoutGivenName rest
where
mkAuxNameWithoutGivenName (rest : List Name) : MetaM (Name × List Name) := do
-- Use a nicer binder name than `[anonymous]`, which can appear in for example `letFun x f` when `f` is not a lambda expression.
-- In this case, we make sure the name is hygienic.
let binderName ← if binderName.isAnonymous then mkFreshUserName `a else pure binderName
if preserveBinderNames then
return (binderName, rest)
else
Expand Down Expand Up @@ -169,11 +183,15 @@ abbrev _root_.Lean.MVarId.intro1P (mvarId : MVarId) : MetaM (FVarId × MVarId) :
abbrev intro1P (mvarId : MVarId) : MetaM (FVarId × MVarId) :=
mvarId.intro1P

private def getIntrosSize : Expr → Nat
private partial def getIntrosSize : Expr → Nat
| .forallE _ _ b _ => getIntrosSize b + 1
| .letE _ _ _ b _ => getIntrosSize b + 1
| .mdata _ b => getIntrosSize b
| _ => 0
| e =>
if let some (_, _, _, b) := e.letFun? then
getIntrosSize b + 1
else
0

/--
Introduce as many binders as possible without unfolding definitions.
Expand Down
26 changes: 26 additions & 0 deletions tests/lean/run/introLetFun.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,26 @@
import Lean.Elab.Tactic.Basic
import Lean.Meta.Tactic.Intro
/-!
# Testing `intro` with `letFun`
-/

/-!
Explicit `intro`.
-/
example : have x := 2; ∀ _ : Nat, x = x := by
intro x _
rfl

/-!
`intros` is aware of `letFun`.
-/
example : have x := 2; ∀ _ : Nat, x = x := by
intros
rfl

/-!
Check that it works for `letFun` with an eta reduced argument.
-/
example (p : Nat → Prop) (h : ∀ x, p x) : letFun 2 p := by
intro
apply h

0 comments on commit 3198109

Please sign in to comment.