Skip to content

Commit 5900f39

Browse files
authored
feat: add [grind intro] attribute (#6888)
This PR adds the `[grind intro]` attribute. It instructs `grind` to mark the introduction rules of an inductive predicate as E-matching theorems.
1 parent b3a8d5b commit 5900f39

File tree

5 files changed

+57
-39
lines changed

5 files changed

+57
-39
lines changed

src/Init/Grind/Tactics.lean

+2-1
Original file line numberDiff line numberDiff line change
@@ -22,7 +22,8 @@ syntax grindFwd := "→ "
2222
syntax grindUsr := &"usr "
2323
syntax grindCases := &"cases "
2424
syntax grindCasesEager := atomic(&"cases" &"eager ")
25-
syntax grindMod := grindEqBoth <|> grindEqRhs <|> grindEq <|> grindEqBwd <|> grindBwd <|> grindFwd <|> grindUsr <|> grindCasesEager <|> grindCases
25+
syntax grindIntro := &"intro "
26+
syntax grindMod := grindEqBoth <|> grindEqRhs <|> grindEq <|> grindEqBwd <|> grindBwd <|> grindFwd <|> grindUsr <|> grindCasesEager <|> grindCases <|> grindIntro
2627
syntax (name := grind) "grind" (grindMod)? : attr
2728
end Attr
2829
end Lean.Parser

src/Lean/Elab/Tactic/Grind.lean

+8-2
Original file line numberDiff line numberDiff line change
@@ -58,7 +58,7 @@ def elabGrindParams (params : Grind.Params) (ps : TSyntaxArray ``Parser.Tactic.
5858
match p with
5959
| `(Parser.Tactic.grindParam| - $id:ident) =>
6060
let declName ← realizeGlobalConstNoOverloadWithInfo id
61-
if (← Grind.isCasesAttrCandidate declName false) then
61+
if let some declName ← Grind.isCasesAttrCandidate? declName false then
6262
Grind.ensureNotBuiltinCases declName
6363
params := { params with casesTypes := (← params.casesTypes.eraseDecl declName) }
6464
else
@@ -82,8 +82,14 @@ def elabGrindParams (params : Grind.Params) (ps : TSyntaxArray ``Parser.Tactic.
8282
| .cases eager =>
8383
withRef p <| Grind.validateCasesAttr declName eager
8484
params := { params with casesTypes := params.casesTypes.insert declName eager }
85+
| .intro =>
86+
if let some info ← Grind.isCasesAttrPredicateCandidate? declName false then
87+
for ctor in info.ctors do
88+
params ← withRef p <| addEMatchTheorem params ctor .default
89+
else
90+
throwError "invalid use of `intro` modifier, `{declName}` is not an inductive predicate"
8591
| .infer =>
86-
if (← Grind.isCasesAttrCandidate declName false) then
92+
if let some declName ← Grind.isCasesAttrCandidate? declName false then
8793
params := { params with casesTypes := params.casesTypes.insert declName false }
8894
if let some info ← isInductivePredicate? declName then
8995
-- If it is an inductive predicate,

src/Lean/Meta/Tactic/Grind/Attr.lean

+9-1
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,7 @@ namespace Lean.Meta.Grind
1212
inductive AttrKind where
1313
| ematch (k : EMatchTheoremKind)
1414
| cases (eager : Bool)
15+
| intro
1516
| infer
1617

1718
/-- Return theorem kind for `stx` of the form `Attr.grindThmMod` -/
@@ -26,6 +27,7 @@ def getAttrKindCore (stx : Syntax) : CoreM AttrKind := do
2627
| `(Parser.Attr.grindMod| usr) => return .ematch .user
2728
| `(Parser.Attr.grindMod| cases) => return .cases false
2829
| `(Parser.Attr.grindMod| cases eager) => return .cases true
30+
| `(Parser.Attr.grindMod| intro) => return .intro
2931
| _ => throwError "unexpected `grind` theorem kind: `{stx}`"
3032

3133
/-- Return theorem kind for `stx` of the form `(Attr.grindMod)?` -/
@@ -64,8 +66,14 @@ builtin_initialize
6466
| .ematch .user => throwInvalidUsrModifier
6567
| .ematch k => addEMatchAttr declName attrKind k
6668
| .cases eager => addCasesAttr declName eager attrKind
69+
| .intro =>
70+
if let some info ← isCasesAttrPredicateCandidate? declName false then
71+
for ctor in info.ctors do
72+
addEMatchAttr ctor attrKind .default
73+
else
74+
throwError "invalid `[grind intro]`, `{declName}` is not an inductive predicate"
6775
| .infer =>
68-
if (← isCasesAttrCandidate declName false) then
76+
if let some declName ← isCasesAttrCandidate? declName false then
6977
addCasesAttr declName false attrKind
7078
if let some info ← isInductivePredicate? declName then
7179
-- If it is an inductive predicate,

src/Lean/Meta/Tactic/Grind/Cases.lean

+12-5
Original file line numberDiff line numberDiff line change
@@ -73,14 +73,21 @@ private def getAlias? (value : Expr) : MetaM (Option Name) :=
7373
else
7474
return none
7575

76-
partial def isCasesAttrCandidate (declName : Name) (eager : Bool) : CoreM Bool := do
76+
partial def isCasesAttrCandidate? (declName : Name) (eager : Bool) : CoreM (Option Name) := do
7777
match (← getConstInfo declName) with
78-
| .inductInfo info => return !info.isRec || !eager
78+
| .inductInfo info => if !info.isRec || !eager then return some declName else return none
7979
| .defnInfo info =>
8080
let some declName ← getAlias? info.value |>.run' {} {}
81-
| return false
82-
isCasesAttrCandidate declName eager
83-
| _ => return false
81+
| return none
82+
isCasesAttrCandidate? declName eager
83+
| _ => return none
84+
85+
def isCasesAttrCandidate (declName : Name) (eager : Bool) : CoreM Bool := do
86+
return (← isCasesAttrCandidate? declName eager).isSome
87+
88+
def isCasesAttrPredicateCandidate? (declName : Name) (eager : Bool) : MetaM (Option InductiveVal) := do
89+
let some declName ← isCasesAttrCandidate? declName eager | return none
90+
isInductivePredicate? declName
8491

8592
def validateCasesAttr (declName : Name) (eager : Bool) : CoreM Unit := do
8693
unless (← isCasesAttrCandidate declName eager) do

tests/lean/run/grind_constProp.lean

+26-30
Original file line numberDiff line numberDiff line change
@@ -80,20 +80,38 @@ def State.get (σ : State) (x : Var) : Val :=
8080
section
8181
attribute [local grind] State.update State.find? State.get State.erase
8282

83-
@[simp, grind =] theorem State.find?_update_self (σ : State) (x : Var) (v : Val) : (σ.update x v).find? x = some v := by
83+
@[simp, grind =] theorem State.find?_nil (x : Var) : find? [] x = none := by
84+
grind
85+
86+
@[simp] theorem State.find?_update_self (σ : State) (x : Var) (v : Val) : (σ.update x v).find? x = some v := by
8487
induction σ, x, v using State.update.induct <;> grind
8588

86-
@[simp, grind =] theorem State.find?_update (σ : State) (v : Val) (h : x ≠ z) : (σ.update x v).find? z = σ.find? z := by
89+
@[simp] theorem State.find?_update (σ : State) (v : Val) (h : x ≠ z) : (σ.update x v).find? z = σ.find? z := by
8790
induction σ, x, v using State.update.induct <;> grind
8891

92+
@[grind =] theorem State.find?_update_eq (σ : State) (v : Val)
93+
: (σ.update x v).find? z = if x = z then some v else σ.find? z := by
94+
grind only [= find?_update_self, = find?_update, cases Or]
95+
8996
@[grind] theorem State.get_of_find? {σ : State} (h : σ.find? x = some v) : σ.get x = v := by
9097
grind
9198

92-
@[simp, grind =] theorem State.find?_erase_self (σ : State) (x : Var) : (σ.erase x).find? x = none := by
99+
@[simp] theorem State.find?_erase_self (σ : State) (x : Var) : (σ.erase x).find? x = none := by
93100
induction σ, x using State.erase.induct <;> grind
94101

95-
@[simp, grind =] theorem State.find?_erase (σ : State) (h : x ≠ z) : (σ.erase x).find? z = σ.find? z := by
102+
@[simp] theorem State.find?_erase (σ : State) (h : x ≠ z) : (σ.erase x).find? z = σ.find? z := by
96103
induction σ, x using State.erase.induct <;> grind
104+
105+
@[simp, grind =] theorem State.find?_erase_eq (σ : State)
106+
: (σ.erase x).find? z = if x = z then none else σ.find? z := by
107+
grind only [= find?_erase_self, = find?_erase, cases Or]
108+
109+
@[grind] theorem State.length_erase_le (σ : State) (x : Var) : (σ.erase x).length ≤ σ.length := by
110+
induction σ, x using erase.induct <;> grind
111+
112+
def State.length_erase_lt (σ : State) (x : Var) : (σ.erase x).length < σ.length.succ := by
113+
grind
114+
97115
end
98116

99117
syntax ident " ↦ " term : term
@@ -206,9 +224,7 @@ def evalExpr (e : Expr) : EvalM Val := do
206224
| c' => .while c' b.simplify
207225

208226
theorem Stmt.simplify_correct (h : (σ, s) ⇓ σ') : (σ, s.simplify) ⇓ σ' := by
209-
-- TODO: we need a mechanism for saying we just want the intro rules
210-
induction h <;> grind [=_ Expr.eval_simplify, Bigstep.skip, Bigstep.assign,
211-
Bigstep.seq, Bigstep.whileFalse, Bigstep.whileTrue, Bigstep.ifTrue, Bigstep.ifFalse]
227+
induction h <;> grind [=_ Expr.eval_simplify, intro Bigstep]
212228

213229
@[simp, grind =] def Expr.constProp (e : Expr) (σ : State) : Expr :=
214230
match e with
@@ -220,13 +236,7 @@ theorem Stmt.simplify_correct (h : (σ, s) ⇓ σ') : (σ, s.simplify) ⇓ σ' :
220236
| una op arg => una op (arg.constProp σ)
221237

222238
@[simp, grind =] theorem Expr.constProp_nil (e : Expr) : e.constProp [] = e := by
223-
induction e <;> grind [State.find?] -- TODO add missing theorem(s) to avoid unfolding `find?`
224-
225-
@[grind] theorem State.length_erase_le (σ : State) (x : Var) : (σ.erase x).length ≤ σ.length := by
226-
induction σ, x using erase.induct <;> grind [State.erase] -- TODO add missing theorem(s)
227-
228-
def State.length_erase_lt (σ : State) (x : Var) : (σ.erase x).length < σ.length.succ := by
229-
grind
239+
induction e <;> grind
230240

231241
@[simp, grind =] def State.join (σ₁ σ₂ : State) : State :=
232242
match σ₁ with
@@ -308,25 +318,11 @@ theorem State.erase_le_of_le_cons (h : σ' ≼ (x, v) :: σ) : σ'.erase x ≼
308318
grind
309319

310320
@[grind] theorem State.erase_le_update (h : σ' ≼ σ) : σ'.erase x ≼ σ.update x v := by
311-
intro y w hf'
312-
-- TODO: can we avoid this hint?
313-
by_cases hxy : x = y <;> grind
321+
grind
314322

315323
@[grind] theorem State.update_le_update (h : σ' ≼ σ) : σ'.update x v ≼ σ.update x v := by
316-
intro y w hf
317-
induction σ generalizing σ' hf with
318-
| nil => grind
319-
| cons zw' σ ih =>
320-
have (z, w') := zw'; simp
321-
have : σ'.erase z ≼ σ := erase_le_of_le_cons h
322-
have ih := ih this
323-
revert ih hf
324-
split <;> simp [*] <;> by_cases hyz : y = z <;> simp (config := { contextual := true }) [*]
325-
next => grind
326-
next => grind
327-
sorry
324+
grind
328325

329-
-- TODO: we are missing theorems here, and cannot seal State functions
330326
@[grind] theorem Expr.eval_constProp_of_sub (e : Expr) (h : σ' ≼ σ) : (e.constProp σ').eval σ = e.eval σ := by
331327
induction e <;> grind
332328

0 commit comments

Comments
 (0)