Skip to content

Commit

Permalink
feat: drop support for termination_by' (leanprover#3033)
Browse files Browse the repository at this point in the history
until around 7fe6881 the way to define well-founded recursions was to
specify a `WellFoundedRelation` on the argument explicitly. This was
rather low-level, for example one had to predict the packing of multiple
arguments into `PProd`s, the packing of mutual functions into `PSum`s,
and the cliques that were calculated.

Then the current `termination_by` syntax was introduced, where you
specify the termination argument at a higher level (one clause per
functions, unpacked arguments), and the `WellFoundedRelation` is found
using type class resolution.

The old syntax was kept around as `termination_by'`. This is not used
anywhere in the lean, std, mathlib or the theorem-proving-in-lean
repositories,
and three occurrences I found in the wild can do without

In particular, it should be possible to express anything that the old
syntax
supported also with the new one, possibly requiring a helper type with a
suitable instance, or the following generic wrapper that now lives in
std
```
def wrap {α : Sort u} {r : α → α → Prop} (h : WellFounded r) (x : α) : {x : α // Acc r x}
```

Since the old syntax is unused, has an unhelpful name and relies on
internals, this removes the support. Now is a good time before the
refactoring that's planned in leanprover#2921.

The test suite was updated without particular surprises.

The parametric `terminationHint` parser is gone, which means we can
match on syntax more easily now, in `expandDecreasingBy?`.
  • Loading branch information
nomeata authored Dec 11, 2023
1 parent 178ab8e commit 5cd90f5
Show file tree
Hide file tree
Showing 36 changed files with 253 additions and 353 deletions.
16 changes: 16 additions & 0 deletions RELEASES.md
Original file line number Diff line number Diff line change
Expand Up @@ -20,6 +20,22 @@ v4.5.0 (development in progress)
```
[PR #2821](https://github.com/leanprover/lean4/pull/2821) and [RFC #2838](https://github.com/leanprover/lean4/issues/2838).

* The low-level `termination_by'` clause is no longer supported.

Migration guide: Use `termination_by` instead, e.g.:
```diff
-termination_by' measure (fun ⟨i, _⟩ => as.size - i)
+termination_by go i _ => as.size - i
```

If the well-founded relation you want to use is not the one that the
`WellFoundedRelation` type class would infer for your termination argument,
you can use `WellFounded.wrap` from the std libarary to explicitly give one:
```diff
-termination_by' ⟨r, hwf⟩
+termination_by _ x => hwf.wrap x
```

v4.4.0
---------

Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Elab/PreDefinition/Main.lean
Original file line number Diff line number Diff line change
Expand Up @@ -100,8 +100,8 @@ def addPreDefinitions (preDefs : Array PreDefinition) (hints : TerminationHints)
let preDefs ← preDefs.mapM ensureNoUnassignedMVarsAtPreDef
let preDefs ← betaReduceLetRecApps preDefs
let cliques := partitionPreDefs preDefs
let mut terminationBy ← liftMacroM <| WF.expandTerminationBy hints.terminationBy? (cliques.map fun ds => ds.map (·.declName))
let mut decreasingBy ← liftMacroM <| WF.expandTerminationHint hints.decreasingBy? (cliques.map fun ds => ds.map (·.declName))
let mut terminationBy ← liftMacroM <| WF.expandTerminationBy? hints.terminationBy? (cliques.map fun ds => ds.map (·.declName))
let mut decreasingBy ← liftMacroM <| WF.expandDecreasingBy? hints.decreasingBy? (cliques.map fun ds => ds.map (·.declName))
let mut hasErrors := false
for preDefs in cliques do
trace[Elab.definition.scc] "{preDefs.map (·.declName)}"
Expand Down
3 changes: 1 addition & 2 deletions src/Lean/Elab/PreDefinition/WF/GuessLex.lean
Original file line number Diff line number Diff line change
Expand Up @@ -32,7 +32,6 @@ In addition to measures derived from `sizeOf xᵢ`, it also considers measures
that assign an order to the functions themselves. This way we can support mutual
function definitions where no arguments decrease from one function to another.
The result of this module is a `TerminationWF`, which is then passed on to `wfRecursion`; this
design is crucial so that whatever we infer in this module could also be written manually by the
user. It would be bad if there are function definitions that can only be processed with the
Expand Down Expand Up @@ -590,7 +589,7 @@ def buildTermWF (declNames : Array Name) (varNamess : Array (Array Name))
declName, vars, body,
implicit := true
}
return .ext termByElements
return termByElements


/--
Expand Down
41 changes: 17 additions & 24 deletions src/Lean/Elab/PreDefinition/WF/Rel.lean
Original file line number Diff line number Diff line change
Expand Up @@ -34,10 +34,10 @@ private partial def unpackUnary (preDef : PreDefinition) (prefixSize : Nat) (mva
let mut varNames ← xs.mapM fun x => x.fvarId!.getUserName
if element.vars.size > varNames.size then
throwErrorAt element.vars[varNames.size]! "too many variable names"
for i in [:element.vars.size] do
let varStx := element.vars[i]!
if varStx.isIdent then
varNames := varNames.set! (varNames.size - element.vars.size + i) varStx.getId
for h : i in [:element.vars.size] do
let varStx := element.vars[i]'h.2
if let `($ident:ident) := varStx then
varNames := varNames.set! (varNames.size - element.vars.size + i) ident.getId
return varNames
let mut mvarId := mvarId
for localDecl in (← Term.getMVarDecl mvarId).lctx, varName in varNames[:prefixSize] do
Expand All @@ -60,25 +60,18 @@ def elabWFRel (preDefs : Array PreDefinition) (unaryPreDefName : Name) (fixedPre
let expectedType := mkApp (mkConst ``WellFoundedRelation [u]) α
trace[Elab.definition.wf] "elabWFRel start: {(← mkFreshTypeMVar).mvarId!}"
withDeclName unaryPreDefName do
match wf with
| TerminationWF.core wfStx =>
let wfRel ← instantiateMVars (← withSynthesize <| elabTermEnsuringType wfStx expectedType)
let pendingMVarIds ← getMVars wfRel
discard <| logUnassignedUsingErrorInfos pendingMVarIds
k wfRel
| TerminationWF.ext elements =>
withRef (getRefFromElems elements) do
let mainMVarId := (← mkFreshExprSyntheticOpaqueMVar expectedType).mvarId!
let [fMVarId, wfRelMVarId, _] ← mainMVarId.apply (← mkConstWithFreshMVarLevels ``invImage) | throwError "failed to apply 'invImage'"
let (d, fMVarId) ← fMVarId.intro1
let subgoals ← unpackMutual preDefs fMVarId d
for (d, mvarId) in subgoals, element in elements, preDef in preDefs do
let mvarId ← unpackUnary preDef fixedPrefixSize mvarId d element
mvarId.withContext do
let value ← Term.withSynthesize <| elabTermEnsuringType element.body (← mvarId.getType)
mvarId.assign value
let wfRelVal ← synthInstance (← inferType (mkMVar wfRelMVarId))
wfRelMVarId.assign wfRelVal
k (← instantiateMVars (mkMVar mainMVarId))
withRef (getRefFromElems wf) do
let mainMVarId := (← mkFreshExprSyntheticOpaqueMVar expectedType).mvarId!
let [fMVarId, wfRelMVarId, _] ← mainMVarId.apply (← mkConstWithFreshMVarLevels ``invImage) | throwError "failed to apply 'invImage'"
let (d, fMVarId) ← fMVarId.intro1
let subgoals ← unpackMutual preDefs fMVarId d
for (d, mvarId) in subgoals, element in wf, preDef in preDefs do
let mvarId ← unpackUnary preDef fixedPrefixSize mvarId d element
mvarId.withContext do
let value ← Term.withSynthesize <| elabTermEnsuringType element.body (← mvarId.getType)
mvarId.assign value
let wfRelVal ← synthInstance (← inferType (mkMVar wfRelMVarId))
wfRelMVarId.assign wfRelVal
k (← instantiateMVars (mkMVar mainMVarId))

end Lean.Elab.WF
175 changes: 74 additions & 101 deletions src/Lean/Elab/PreDefinition/WF/TerminationHint.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,114 +5,112 @@ Authors: Leonardo de Moura
-/
import Lean.Parser.Command

set_option autoImplicit false

namespace Lean.Elab.WF

/-! # Support for `decreasing_by` and `termination_by'` notations -/
/-! # Support for `decreasing_by` -/

structure TerminationHintValue where
structure DecreasingByTactic where
ref : Syntax
value : Syntax
value : Lean.TSyntax `Lean.Parser.Tactic.tacticSeq
deriving Inhabited

inductive TerminationHint where
inductive DecreasingBy where
| none
| one (val : TerminationHintValue)
| many (map : NameMap TerminationHintValue)
| one (val : DecreasingByTactic)
| many (map : NameMap DecreasingByTactic)
deriving Inhabited

open Parser.Command in
/--
This function takes a user-specified `decreasing_by` or `termination_by'` clause (as `Sytnax`).
If it is a `terminathionHintMany` (a set of clauses guarded by the function name) then
This function takes a user-specified `decreasing_by` clause (as `Sytnax`).
If it is a `decreasingByMany` (a set of clauses guarded by the function name) then
* checks that all mentioned names exist in the current declaration
* check that at most one function from each clique is mentioned
and sort the entries by function name.
-/
def expandTerminationHint (terminationHint? : Option Syntax) (cliques : Array (Array Name)) : MacroM TerminationHint := do
let some terminationHint := terminationHint? | return TerminationHint.none
let ref := terminationHint
match terminationHint with
| `(terminationByCore|termination_by' $hint1:terminationHint1)
| `(decreasingBy|decreasing_by $hint1:terminationHint1) =>
return TerminationHint.one { ref, value := hint1.raw[0] }
| `(terminationByCore|termination_by' $hints:terminationHintMany)
| `(decreasingBy|decreasing_by $hints:terminationHintMany) => do
let m ← hints.raw[0].getArgs.foldlM (init := {}) fun m arg =>
def expandDecreasingBy? (decreasingBy? : Option Syntax) (cliques : Array (Array Name)) : MacroM DecreasingBy := do
let some decreasingBy := decreasingBy? | return DecreasingBy.none
let ref := decreasingBy
match decreasingBy with
| `(decreasingBy|decreasing_by $hint1:tacticSeq) =>
return DecreasingBy.one { ref, value := hint1 }
| `(decreasingBy|decreasing_by $hints:decreasingByMany) => do
let m ← hints.raw[0].getArgs.foldlM (init := {}) fun m arg => do
let arg : TSyntax `decreasingByElement := ⟨arg⟩ -- cannot use syntax pattern match with lookahead
let `(decreasingByElement| $declId:ident => $tac:tacticSeq) := arg | Macro.throwUnsupported
let declName? := cliques.findSome? fun clique => clique.findSome? fun declName =>
if arg[0].getId.isSuffixOf declName then some declName else none
if declId.getId.isSuffixOf declName then some declName else none
match declName? with
| none => Macro.throwErrorAt arg[0] s!"function '{arg[0].getId}' not found in current declaration"
| some declName => return m.insert declName { ref := arg, value := arg[2] }
| none => Macro.throwErrorAt declId s!"function '{declId.getId}' not found in current declaration"
| some declName => return m.insert declName { ref := arg, value := tac }
for clique in cliques do
let mut found? := Option.none
for declName in clique do
if let some { ref, .. } := m.find? declName then
if let some found := found? then
Macro.throwErrorAt ref s!"invalid termination hint element, '{declName}' and '{found}' are in the same clique"
found? := some declName
return TerminationHint.many m
return DecreasingBy.many m
| _ => Macro.throwUnsupported

def TerminationHint.markAsUsed (t : TerminationHint) (clique : Array Name) : TerminationHint :=
def DecreasingBy.markAsUsed (t : DecreasingBy) (clique : Array Name) : DecreasingBy :=
match t with
| TerminationHint.none => TerminationHint.none
| TerminationHint.one .. => TerminationHint.none
| TerminationHint.many m => Id.run do
| DecreasingBy.none => DecreasingBy.none
| DecreasingBy.one .. => DecreasingBy.none
| DecreasingBy.many m => Id.run do
for declName in clique do
if m.contains declName then
let m := m.erase declName
if m.isEmpty then
return TerminationHint.none
return DecreasingBy.none
else
return TerminationHint.many m
return DecreasingBy.many m
return t

def TerminationHint.find? (t : TerminationHint) (clique : Array Name) : Option TerminationHintValue :=
def DecreasingBy.find? (t : DecreasingBy) (clique : Array Name) : Option DecreasingByTactic :=
match t with
| TerminationHint.none => Option.none
| TerminationHint.one v => some v
| TerminationHint.many m => clique.findSome? m.find?
| DecreasingBy.none => Option.none
| DecreasingBy.one v => some v
| DecreasingBy.many m => clique.findSome? m.find?

def TerminationHint.ensureAllUsed (t : TerminationHint) : MacroM Unit := do
def DecreasingBy.ensureAllUsed (t : DecreasingBy) : MacroM Unit := do
match t with
| TerminationHint.one v => Macro.throwErrorAt v.ref "unused termination hint element"
| TerminationHint.many m => m.forM fun _ v => Macro.throwErrorAt v.ref "unused termination hint element"
| DecreasingBy.one v => Macro.throwErrorAt v.ref "unused termination hint element"
| DecreasingBy.many m => m.forM fun _ v => Macro.throwErrorAt v.ref "unused termination hint element"
| _ => pure ()

/-! # Support for `termination_by` and `termination_by'` notation -/
/-! # Support for `termination_by` notation -/

/-- A single `termination_by` clause -/
structure TerminationByElement where
ref : Syntax
declName : Name
vars : Array Syntax
body : Syntax
vars : TSyntaxArray [`ident, ``Lean.Parser.Term.hole]
body : Term
implicit : Bool
deriving Inhabited

/-- `terminatoin_by` clauses, grouped by clique -/
/-- `termination_by` clauses, grouped by clique -/
structure TerminationByClique where
elements : Array TerminationByElement
used : Bool := false

/--
A `termination_by'` or `termination_by` hint, as specified by the user
A `termination_by` hint, as specified by the user
-/
inductive TerminationBy where
/-- `termination_by'` -/
| core (hint : TerminationHint)
/-- `termination_by` -/
| ext (cliques : Array TerminationByClique)
structure TerminationBy where
cliques : Array TerminationByClique
deriving Inhabited

/--
A `termination_by'` or `termination_by` hint, as applicable to a single clique
A `termination_by` hint, as applicable to a single clique
-/
inductive TerminationWF where
| core (stx : Syntax)
| ext (clique : Array TerminationByElement)
abbrev TerminationWF := Array TerminationByElement

/-
open Parser.Command in
/--
Expands the syntax for a `termination_by` clause, checking that
* each function is mentioned once
* each function mentioned actually occurs in the current declaration
Expand All @@ -125,8 +123,9 @@ def terminationByElement := leading_parser ppLine >> (ident <|> hole) >> many
def terminationBy := leading_parser ppLine >> "termination_by " >> many1chIndent terminationByElement
```
-/
open Parser.Command in
private def expandTerminationByNonCore (hint : Syntax) (cliques : Array (Array Name)) : MacroM TerminationBy := do
def expandTerminationBy? (hint? : Option Syntax) (cliques : Array (Array Name)) :
MacroM TerminationBy := do
let some hint := hint? | return { cliques := #[] }
let `(terminationBy|termination_by $elementStxs*) := hint | Macro.throwUnsupported
let mut alreadyFound : NameSet := {}
let mut elseElemStx? := none
Expand Down Expand Up @@ -170,71 +169,45 @@ private def expandTerminationByNonCore (hint : Syntax) (cliques : Array (Array N
result := result.push { elements }
if !usedElse && elseElemStx?.isSome then
withRef elseElemStx?.get! <| Macro.throwError s!"invalid `termination_by` syntax, unnecessary else-case"
return TerminationBy.ext result

/--
Expands the syntax for a `termination_by` or `termination_by'` clause, using the appropriate function
-/
def expandTerminationBy (hint? : Option Syntax) (cliques : Array (Array Name)) : MacroM TerminationBy :=
if let some hint := hint? then
if hint.isOfKind ``Parser.Command.terminationByCore then
return TerminationBy.core (← expandTerminationHint hint? cliques)
else if hint.isOfKind ``Parser.Command.terminationBy then
expandTerminationByNonCore hint cliques
else
Macro.throwUnsupported
else
return TerminationBy.core TerminationHint.none
return ⟨result⟩

open Parser.Command in
def TerminationWF.unexpand : TerminationWF → MetaM Syntax
| .ext elements => do
let elementStxs ← elements.mapM fun element => do
let fn : Ident := mkIdent (← unresolveNameGlobal element.declName)
let body : Term := ⟨element.body⟩
let vars : Array Ident := element.vars.map TSyntax.mk
`(terminationByElement|$fn $vars* => $body)
`(terminationBy|termination_by $elementStxs*)
| .core _ => unreachable! -- we don't synthetize termination_by' syntax
def TerminationWF.unexpand (elements : TerminationWF) : MetaM Syntax := do
let elementStxs ← elements.mapM fun element => do
let fn : Ident := mkIdent (← unresolveNameGlobal element.declName)
`(terminationByElement|$fn $element.vars* => $element.body)
`(terminationBy|termination_by $elementStxs*)

def TerminationBy.markAsUsed (t : TerminationBy) (cliqueNames : Array Name) : TerminationBy :=
match t with
| core hint => core (hint.markAsUsed cliqueNames)
| ext cliques => ext <| cliques.map fun clique =>
.mk <| t.cliques.map fun clique =>
if cliqueNames.any fun name => clique.elements.any fun elem => elem.declName == name then
{ clique with used := true }
else
clique

def TerminationBy.find? (t : TerminationBy) (cliqueNames : Array Name) : Option TerminationWF :=
match t with
| core hint => hint.find? cliqueNames |>.map fun v => TerminationWF.core v.value
| ext cliques =>
cliques.findSome? fun clique =>
if cliqueNames.any fun name => clique.elements.any fun elem => elem.declName == name then
some <| TerminationWF.ext clique.elements
else
none
t.cliques.findSome? fun clique =>
if cliqueNames.any fun name => clique.elements.any fun elem => elem.declName == name then
some <| clique.elements
else
none

def TerminationByClique.allImplicit (c : TerminationByClique) : Bool :=
c.elements.all fun elem => elem.implicit

def TerminationByClique.getExplicitElement? (c : TerminationByClique) : Option TerminationByElement :=
c.elements.find? (!·.implicit)

def TerminationBy.ensureAllUsed (t : TerminationBy) : MacroM Unit :=
match t with
| core hint => hint.ensureAllUsed
| ext cliques => do
let hasUsedAllImplicit := cliques.any fun c => c.allImplicit && c.used
let mut reportedAllImplicit := true
for clique in cliques do
unless clique.used do
if let some explicitElem := clique.getExplicitElement? then
Macro.throwErrorAt explicitElem.ref "unused termination hint element"
else if !hasUsedAllImplicit then
unless reportedAllImplicit do
reportedAllImplicit := true
Macro.throwErrorAt clique.elements[0]!.ref "unused termination hint element"
def TerminationBy.ensureAllUsed (t : TerminationBy) : MacroM Unit := do
let hasUsedAllImplicit := t.cliques.any fun c => c.allImplicit && c.used
let mut reportedAllImplicit := true
for clique in t.cliques do
unless clique.used do
if let some explicitElem := clique.getExplicitElement? then
Macro.throwErrorAt explicitElem.ref "unused termination hint element"
else if !hasUsedAllImplicit then
unless reportedAllImplicit do
reportedAllImplicit := true
Macro.throwErrorAt clique.elements[0]!.ref "unused termination hint element"

end Lean.Elab.WF
Loading

0 comments on commit 5cd90f5

Please sign in to comment.