Skip to content

Commit

Permalink
update monomorphic testing
Browse files Browse the repository at this point in the history
  • Loading branch information
PratherConid committed May 2, 2024
1 parent 17a4e94 commit bde829b
Show file tree
Hide file tree
Showing 2 changed files with 25 additions and 18 deletions.
40 changes: 23 additions & 17 deletions Auto/Lib/ExprExtra.lean
Original file line number Diff line number Diff line change
Expand Up @@ -214,25 +214,31 @@ def Expr.numLeadingDepArgs : Expr → Nat
0
| _ => 0

/-- Turn all `Prop` binders into `True` -/
private partial def Expr.isMonomorphicFactAux : Expr → MetaM Expr
| .forallE name ty body bi => do
let ty := if (← Meta.isProp ty) ∧ !body.hasLooseBVars then .const ``False [] else ty
Meta.withLocalDecl name bi ty fun x => do
let body := body.instantiate1 x
let body ← isMonomorphicFactAux body
Meta.mkForallFVars #[x] body
| _ => pure (.const ``False [])

/--
Test whether `e` is a monomorphic fact.
`e` is a monomorphic fact iff for all subterms `t : α` of `e`
where `α` is not of type `Prop`, `α` does not depend on bound
variables.
Check whether the leading `∀` quantifiers of expression `e`
violates the quasi-monomorphic condition
-/
def Expr.isMonomorphicFact (e : Expr) : MetaM Bool := do
let e ← Expr.isMonomorphicFactAux e
return (Expr.depArgs e).size == 0
partial def Expr.leadingForallQuasiMonomorphicAux (fvars : Array FVarId) (e : Expr) : MetaM Bool :=
match e with
| .forallE name ty body bi => Meta.withLocalDecl name bi ty fun x => do
let Expr.fvar xid := x
| throwError "Expr.leadingForallQuasiMonomorphic :: Unexpected error"
let bodyi := body.instantiate1 x
if ← Meta.isProp ty then
if !(← Meta.isProp bodyi) then
return false
if body.hasLooseBVar 0 then
return false
return (← Expr.leadingForallQuasiMonomorphicAux fvars ty) &&
(← Expr.leadingForallQuasiMonomorphicAux (fvars.push xid) bodyi)
else
let fvarSet := HashSet.empty.insertMany fvars
if ty.hasAnyFVar fvarSet.contains then
return false
Expr.leadingForallQuasiMonomorphicAux (fvars.push xid) bodyi
| _ => return true

def Expr.leadingForallQuasiMonomorphic := Expr.leadingForallQuasiMonomorphicAux #[]

/--
This should only be used when we're sure that reducing `ty`
Expand Down
3 changes: 2 additions & 1 deletion Auto/Translation/Monomorphization.lean
Original file line number Diff line number Diff line change
Expand Up @@ -386,7 +386,8 @@ def LemmaInst.matchConstInst (ci : ConstInst) (li : LemmaInst) : MetaM (HashSet
def LemmaInst.monomorphic? (li : LemmaInst) : MetaM (Option LemmaInst) := do
if li.params.size != 0 then
return .none
if !(← Expr.isMonomorphicFact li.type) then
let li := {li with type := ← prepReduceExpr li.type}
if !(← Expr.leadingForallQuasiMonomorphic li.type) then
return .none
Meta.withNewMCtxDepth do
let (_, mvars, mi) ← MLemmaInst.ofLemmaInst li
Expand Down

0 comments on commit bde829b

Please sign in to comment.