Skip to content

Commit

Permalink
feat: functional induction (#3432)
Browse files Browse the repository at this point in the history
This adds the concept of **functional induction** to lean.

Derived from the definition of a (possibly mutually) recursive function,
a **functional
induction principle** is tailored to proofs about that function. For
example from:

```
def ackermann : Nat → Nat → Nat
  | 0, m => m + 1
  | n+1, 0 => ackermann n 1
  | n+1, m+1 => ackermann n (ackermann (n + 1) m)
derive_functional_induction ackermann
```
we get
```
ackermann.induct (motive : Nat → Nat → Prop) (case1 : ∀ (m : Nat), motive 0 m)
  (case2 : ∀ (n : Nat), motive n 1 → motive (Nat.succ n) 0)
  (case3 : ∀ (n m : Nat), motive (n + 1) m → motive n (ackermann (n + 1) m) → motive (Nat.succ n) (Nat.succ m))
  (x x : Nat) : motive x x
```

At the moment, the user has to ask for the functional induction
principle explicitly using
```
derive_functional_induction ackermann
```

The module docstring of `Lean/Meta/Tactic/FunInd.lean` contains more
details on the
design and implementation of this command.

More convenience around this (e.g. a `functional induction` tactic) will
follow eventually.


This PR includes a bunch of `PSum`/`PSigma` related functions in the
`Lean.Tactic.FunInd`
namespace. I plan to move these to `PackArgs`/`PackMutual` afterwards,
and do some cleaning
up as I do that.

---------

Co-authored-by: David Thrane Christiansen <[email protected]>
Co-authored-by: Leonardo de Moura <[email protected]>
  • Loading branch information
3 people authored Mar 5, 2024
1 parent ce77518 commit 8038604
Show file tree
Hide file tree
Showing 14 changed files with 2,347 additions and 4 deletions.
3 changes: 2 additions & 1 deletion .github/workflows/ci.yml
Original file line number Diff line number Diff line change
Expand Up @@ -140,7 +140,8 @@ jobs:
"shell": "msys2 {0}",
"CMAKE_OPTIONS": "-G \"Unix Makefiles\" -DUSE_GMP=OFF",
// for reasons unknown, interactivetests are flaky on Windows
"CTEST_OPTIONS": "--repeat until-pass:2",
// also, the liasolver test hits “too many exported symbols”
"CTEST_OPTIONS": "--repeat until-pass:2 -E 'leanbenchtest_liasolver.lean'",
"llvm-url": "https://github.com/leanprover/lean-llvm/releases/download/15.0.1/lean-llvm-x86_64-w64-windows-gnu.tar.zst",
"prepare-llvm": "../script/prepare-llvm-mingw.sh lean-llvm*",
"binary-check": "ldd"
Expand Down
20 changes: 20 additions & 0 deletions RELEASES.md
Original file line number Diff line number Diff line change
Expand Up @@ -11,6 +11,26 @@ of each version.
v4.8.0 (development in progress)
---------

* New command `derive_functinal_induction`:

Derived from the definition of a (possibly mutually) recursive function
defined by well-founded recursion, a **functional induction principle** is
tailored to proofs about that function. For example from:
```
def ackermann : Nat → Nat → Nat
| 0, m => m + 1
| n+1, 0 => ackermann n 1
| n+1, m+1 => ackermann n (ackermann (n + 1) m)
derive_functional_induction ackermann
```
we get
```
ackermann.induct (motive : Nat → Nat → Prop) (case1 : ∀ (m : Nat), motive 0 m)
(case2 : ∀ (n : Nat), motive n 1 → motive (Nat.succ n) 0)
(case3 : ∀ (n m : Nat), motive (n + 1) m → motive n (ackermann (n + 1) m) → motive (Nat.succ n) (Nat.succ m))
(x x : Nat) : motive x x
```

v4.7.0
---------

Expand Down
3 changes: 3 additions & 0 deletions src/Lean/CoreM.lean
Original file line number Diff line number Diff line change
Expand Up @@ -289,6 +289,9 @@ def Exception.isMaxHeartbeat (ex : Exception) : Bool :=
def mkArrow (d b : Expr) : CoreM Expr :=
return Lean.mkForall (← mkFreshUserName `x) BinderInfo.default d b

/-- Iterated `mkArrow`, creates the expression `a₁ → a₂ → … → aₙ → b`. Also see `arrowDomainsN`. -/
def mkArrowN (ds : Array Expr) (e : Expr) : CoreM Expr := ds.foldrM mkArrow e

def addDecl (decl : Declaration) : CoreM Unit := do
profileitM Exception "type checking" (← getOptions) do
withTraceNode `Kernel (fun _ => return m!"typechecking declaration") do
Expand Down
10 changes: 10 additions & 0 deletions src/Lean/Meta/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1347,6 +1347,16 @@ private def withNewMCtxDepthImp (allowLevelAssignments : Bool) (x : MetaM α) :
finally
modify fun s => { s with mctx := saved.mctx, postponed := saved.postponed }

/--
Removes `fvarId` from the local context, and replaces occurrences of it with `e`.
It is the responsibility of the caller to ensure that `e` is well-typed in the context
of any occurrence of `fvarId`.
-/
def withReplaceFVarId {α} (fvarId : FVarId) (e : Expr) : MetaM α → MetaM α :=
withReader fun ctx => { ctx with
lctx := ctx.lctx.replaceFVarId fvarId e
localInstances := ctx.localInstances.erase fvarId }

/--
`withNewMCtxDepth k` executes `k` with a higher metavariable context depth,
where metavariables created outside the `withNewMCtxDepth` (with a lower depth) cannot be assigned.
Expand Down
12 changes: 9 additions & 3 deletions src/Lean/Meta/Match/MatchEqs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -110,6 +110,8 @@ def unfoldNamedPattern (e : Expr) : MetaM Expr := do
- `type` is the resulting type for `altType`.
We use the `mask` to build the splitter proof. See `mkSplitterProof`.
This can be used to use the alternative of a match expression in its splitter.
-/
partial def forallAltTelescope (altType : Expr) (altNumParams numDiscrEqs : Nat)
(k : (ys : Array Expr) → (eqs : Array Expr) → (args : Array Expr) → (mask : Array Bool) → (type : Expr) → MetaM α)
Expand All @@ -132,9 +134,11 @@ where
let some k := args.getIdx? lhs | unreachable!
let mask := mask.set! k false
let args := args.map fun arg => if arg == lhs then rhs else arg
let args := args.push (← mkEqRefl rhs)
let arg ← mkEqRefl rhs
let typeNew := typeNew.replaceFVar lhs rhs
return (← go ys eqs args (mask.push false) (i+1) typeNew)
return ← withReplaceFVarId lhs.fvarId! rhs do
withReplaceFVarId y.fvarId! arg do
go ys eqs (args.push arg) (mask.push false) (i+1) typeNew
go (ys.push y) eqs (args.push y) (mask.push true) (i+1) typeNew
else
let arg ← if let some (_, _, rhs) ← matchEq? d then
Expand All @@ -152,7 +156,9 @@ where
they are not eagerly evaluated. -/
if ys.size == 1 then
if (← inferType ys[0]!).isConstOf ``Unit && !(← dependsOn type ys[0]!.fvarId!) then
return (← k #[] #[] #[mkConst ``Unit.unit] #[false] type)
let rhs := mkConst ``Unit.unit
return ← withReplaceFVarId ys[0]!.fvarId! rhs do
return (← k #[] #[] #[rhs] #[false] type)
k ys eqs args mask type

isNamedPatternProof (type : Expr) (h : Expr) : Bool :=
Expand Down
255 changes: 255 additions & 0 deletions src/Lean/Meta/Match/MatcherApp/Transform.lean
Original file line number Diff line number Diff line change
Expand Up @@ -156,4 +156,259 @@ def refineThrough? (matcherApp : MatcherApp) (e : Expr) :
catch _ =>
return none


/--
Given `n` and a non-dependent function type `α₁ → α₂ → ... → αₙ → Sort u`, returns the
types `α₁, α₂, ..., αₙ`. Throws an error if there are not at least `n` argument types or if a
later argument type depends on a prior one (i.e., it's a dependent function type).
This can be used to infer the expected type of the alternatives when constructing a `MatcherApp`.
-/
-- TODO: Which is the natural module for this?
def arrowDomainsN (n : Nat) (type : Expr) : MetaM (Array Expr) := do
let mut type := type
let mut ts := #[]
for i in [:n] do
type ← whnfForall type
let Expr.forallE _ α β _ ← pure type | throwError "expected {n} arguments, got {i}"
if β.hasLooseBVars then throwError "unexpected dependent type"
ts := ts.push α
type := β
return ts

/--
Performs a possibly type-changing transformation to a `MatcherApp`.
* `onParams` is run on each parameter and discriminant
* `onMotive` runs on the body of the motive, and is passed the motive parameters
(one for each `MatcherApp.discrs`)
* `onAlt` runs on each alternative, and is passed the expected type of the alternative,
as inferred from the motive
* `onRemaining` runs on the remaining arguments (and may change their number)
If `useSplitter` is true, the matcher is replaced with the splitter.
NB: Not all operations on `MatcherApp` can handle one `matcherName` is a splitter.
The array `addEqualities`, if provided, indicates for which of the discriminants an equality
connecting the discriminant to the parameters of the alternative (like in `match h : x with …`)
should be added (if it is isn't already there).
This function works even if the the type of alternatives do *not* fit the inferred type. This
allows you to post-process the `MatcherApp` with `MatcherApp.inferMatchType`, which will
infer a type, given all the alternatives.
-/
def transform (matcherApp : MatcherApp)
(useSplitter := false)
(addEqualities : Array Bool := mkArray matcherApp.discrs.size false)
(onParams : Expr → MetaM Expr := pure)
(onMotive : Array Expr → Expr → MetaM Expr := fun _ e => pure e)
(onAlt : Expr → Expr → MetaM Expr := fun _ e => pure e)
(onRemaining : Array Expr → MetaM (Array Expr) := pure) :
MetaM MatcherApp := do

if addEqualities.size != matcherApp.discrs.size then
throwError "MatcherApp.transform: addEqualities has wrong size"

-- Do not add equalities when the matcher already does so
let addEqualities := Array.zipWith addEqualities matcherApp.discrInfos fun b di =>
if di.hName?.isSome then false else b

-- We also handle CasesOn applications here, and need to treat them specially in a
-- few places.
-- TODO: Expand MatcherApp with the necessary fields to make this more uniform
-- (in particular, include discrEq and whether there is a splitter)
let isCasesOn := isCasesOnRecursor (← getEnv) matcherApp.matcherName

let numDiscrEqs ←
if isCasesOn then pure 0 else
match ← getMatcherInfo? matcherApp.matcherName with
| some info => pure info.getNumDiscrEqs
| none => throwError "matcher {matcherApp.matcherName} has no MatchInfo found"

let params' ← matcherApp.params.mapM onParams
let discrs' ← matcherApp.discrs.mapM onParams


let (motive', uElim) ← lambdaTelescope matcherApp.motive fun motiveArgs motiveBody => do
unless motiveArgs.size == matcherApp.discrs.size do
throwError "unexpected matcher application, motive must be lambda expression with #{matcherApp.discrs.size} arguments"
let mut motiveBody' ← onMotive motiveArgs motiveBody

-- Prepend (x = e) → to the motive when an equality is requested
for arg in motiveArgs, discr in discrs', b in addEqualities do if b then
motiveBody' ← mkArrow (← mkEq discr arg) motiveBody'

return (← mkLambdaFVars motiveArgs motiveBody', ← getLevel motiveBody')

let matcherLevels ← match matcherApp.uElimPos? with
| none => pure matcherApp.matcherLevels
| some pos => pure <| matcherApp.matcherLevels.set! pos uElim

-- We pass `Eq.refl`s for all the equations we added as extra arguments
-- (and count them along the way)
let mut remaining' := #[]
let mut extraEqualities : Nat := 0
for discr in discrs'.reverse, b in addEqualities.reverse do if b then
remaining' := remaining'.push (← mkEqRefl discr)
extraEqualities := extraEqualities + 1

if useSplitter && !isCasesOn then
-- We replace the matcher with the splitter
let matchEqns ← Match.getEquationsFor matcherApp.matcherName
let splitter := matchEqns.splitterName

let aux1 := mkAppN (mkConst matcherApp.matcherName matcherLevels.toList) params'
let aux1 := mkApp aux1 motive'
let aux1 := mkAppN aux1 discrs'
unless (← isTypeCorrect aux1) do
logError m!"failed to transform matcher, type error when constructing new motive:{indentExpr aux1}"
check aux1
let origAltTypes ← arrowDomainsN matcherApp.alts.size (← inferType aux1)

let aux2 := mkAppN (mkConst splitter matcherLevels.toList) params'
let aux2 := mkApp aux2 motive'
let aux2 := mkAppN aux2 discrs'
unless (← isTypeCorrect aux2) do
logError m!"failed to transform matcher, type error when constructing new motive:{indentExpr aux2}"
check aux2
let altTypes ← arrowDomainsN matcherApp.alts.size (← inferType aux2)

let mut alts' := #[]
for alt in matcherApp.alts,
numParams in matcherApp.altNumParams,
splitterNumParams in matchEqns.splitterAltNumParams,
origAltType in origAltTypes,
altType in altTypes do
let alt' ← Match.forallAltTelescope origAltType (numParams - numDiscrEqs) 0 fun ys _eqs args _mask _bodyType => do
let altType ← instantiateForall altType ys
-- The splitter inserts its extra paramters after the first ys.size parameters, before
-- the parameters for the numDiscrEqs
forallBoundedTelescope altType (splitterNumParams - ys.size) fun ys2 altType => do
forallBoundedTelescope altType numDiscrEqs fun ys3 altType => do
forallBoundedTelescope altType extraEqualities fun ys4 altType => do
let alt ← try instantiateLambda alt (args ++ ys3)
catch _ => throwError "unexpected matcher application, insufficient number of parameters in alternative"
let alt' ← onAlt altType alt
mkLambdaFVars (ys ++ ys2 ++ ys3 ++ ys4) alt'
alts' := alts'.push alt'

remaining' := remaining' ++ (← onRemaining matcherApp.remaining)

return { matcherApp with
matcherName := splitter
matcherLevels := matcherLevels
params := params'
motive := motive'
discrs := discrs'
altNumParams := matchEqns.splitterAltNumParams.map (· + extraEqualities)
alts := alts'
remaining := remaining'
}
else
let aux := mkAppN (mkConst matcherApp.matcherName matcherLevels.toList) params'
let aux := mkApp aux motive'
let aux := mkAppN aux discrs'
unless (← isTypeCorrect aux) do
-- check aux
logError m!"failed to transform matcher, type error when constructing new motive:{indentExpr aux}"
check aux
let altTypes ← arrowDomainsN matcherApp.alts.size (← inferType aux)

let mut alts' := #[]
for alt in matcherApp.alts,
numParams in matcherApp.altNumParams,
altType in altTypes do
let alt' ← forallBoundedTelescope altType numParams fun xs altType => do
forallBoundedTelescope altType extraEqualities fun ys4 altType => do
let alt ← instantiateLambda alt xs
let alt' ← onAlt altType alt
mkLambdaFVars (xs ++ ys4) alt'
alts' := alts'.push alt'

remaining' := remaining' ++ (← onRemaining matcherApp.remaining)

return { matcherApp with
matcherLevels := matcherLevels
params := params'
motive := motive'
discrs := discrs'
altNumParams := matcherApp.altNumParams.map (· + extraEqualities)
alts := alts'
remaining := remaining'
}



/--
Given a `MatcherApp`, replaces the motive with one that is inferred from the actual types of the
alternatives.
For example, given
```
(match (motive := Nat → Unit → ?) n with
0 => 1
_ => true) ()
```
(for any `?`; the motive’s result type be ignored) will give this type
```
(match n with
| 0 => Nat
| _ => Bool)
```
The given `MatcherApp` must not use a splitter in `matcherName`.
The resulting expression *will* use the splitter corresponding to `matcherName` (this is necessary
for the construction).
Interally, this needs to reduce the matcher in a given branch; this is done using
`Split.simpMatchTarget`.
-/
def inferMatchType (matcherApp : MatcherApp) : MetaM MatcherApp := do
-- In matcherApp.motive, replace the (dummy) matcher body with a type
-- derived from the inferred types of the alterantives
let nExtra := matcherApp.remaining.size
matcherApp.transform (useSplitter := true)
(onMotive := fun motiveArgs body => do
let extraParams ← arrowDomainsN nExtra body
let propMotive ← mkLambdaFVars motiveArgs (.sort levelZero)
let propAlts ← matcherApp.alts.mapM fun termAlt =>
lambdaTelescope termAlt fun xs termAltBody => do
-- We have alt parameters and parameters corresponding to the extra args
let xs1 := xs[0 : xs.size - nExtra]
let xs2 := xs[xs.size - nExtra : xs.size]
-- logInfo m!"altIH: {xs} => {altIH}"
let altType ← inferType termAltBody
for x in xs2 do
if altType.hasAnyFVar (· == x.fvarId!) then
throwError "Type {altType} of alternative {termAlt} still depends on {x}"
-- logInfo m!"altIH type: {altType}"
mkLambdaFVars xs1 altType
let matcherLevels ← match matcherApp.uElimPos? with
| none => pure matcherApp.matcherLevels
| some pos => pure <| matcherApp.matcherLevels.set! pos levelOne
let typeMatcherApp := { matcherApp with
motive := propMotive
matcherLevels := matcherLevels
discrs := motiveArgs
alts := propAlts
remaining := #[]
}
mkArrowN extraParams typeMatcherApp.toExpr
)
(onAlt := fun expAltType alt => do
let altType ← inferType alt
let eq ← mkEq expAltType altType
let proof ← mkFreshExprSyntheticOpaqueMVar eq
let goal := proof.mvarId!
-- logInfo m!"Goal: {goal}"
let goal ← Split.simpMatchTarget goal
-- logInfo m!"Goal after splitting: {goal}"
try
goal.refl
catch _ =>
logInfo m!"Cannot close goal after splitting: {goal}"
goal.admit
mkEqMPR proof alt
)

end Lean.Meta.MatcherApp
1 change: 1 addition & 0 deletions src/Lean/Meta/Tactic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -37,3 +37,4 @@ import Lean.Meta.Tactic.IndependentOf
import Lean.Meta.Tactic.Symm
import Lean.Meta.Tactic.Backtrack
import Lean.Meta.Tactic.SolveByElim
import Lean.Meta.Tactic.FunInd
Loading

0 comments on commit 8038604

Please sign in to comment.