-
Notifications
You must be signed in to change notification settings - Fork 546
/
Copy pathTactics.lean
91 lines (80 loc) · 3.44 KB
/
Tactics.lean
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
/-
Copyright (c) 2024 Amazon.com, Inc. or its affiliates. All Rights Reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
-/
prelude
import Init.Tactics
namespace Lean.Parser
/--
Reset all `grind` attributes. This command is intended for testing purposes only and should not be used in applications.
-/
syntax (name := resetGrindAttrs) "%reset_grind_attrs" : command
namespace Attr
syntax grindEq := "= "
syntax grindEqBoth := atomic("_" "=" "_ ")
syntax grindEqRhs := atomic("=" "_ ")
syntax grindEqBwd := atomic("←" "= ")
syntax grindBwd := "← "
syntax grindFwd := "→ "
syntax grindUsr := &"usr "
syntax grindCases := &"cases "
syntax grindCasesEager := atomic(&"cases" &"eager ")
syntax grindIntro := &"intro "
syntax grindMod := grindEqBoth <|> grindEqRhs <|> grindEq <|> grindEqBwd <|> grindBwd <|> grindFwd <|> grindUsr <|> grindCasesEager <|> grindCases <|> grindIntro
syntax (name := grind) "grind" (grindMod)? : attr
end Attr
end Lean.Parser
namespace Lean.Grind
/--
The configuration for `grind`.
Passed to `grind` using, for example, the `grind (config := { matchEqs := true })` syntax.
-/
structure Config where
/-- If `trace` is `true`, `grind` records used E-matching theorems and case-splits. -/
trace : Bool := false
/-- Maximum number of case-splits in a proof search branch. It does not include splits performed during normalization. -/
splits : Nat := 8
/-- Maximum number of E-matching (aka heuristic theorem instantiation) rounds before each case split. -/
ematch : Nat := 5
/--
Maximum term generation.
The input goal terms have generation 0. When we instantiate a theorem using a term from generation `n`,
the new terms have generation `n+1`. Thus, this parameter limits the length of an instantiation chain. -/
gen : Nat := 5
/-- Maximum number of theorem instances generated using E-matching in a proof search tree branch. -/
instances : Nat := 1000
/-- If `matchEqs` is `true`, `grind` uses `match`-equations as E-matching theorems. -/
matchEqs : Bool := true
/-- If `splitMatch` is `true`, `grind` performs case-splitting on `match`-expressions during the search. -/
splitMatch : Bool := true
/-- If `splitIte` is `true`, `grind` performs case-splitting on `if-then-else` expressions during the search. -/
splitIte : Bool := true
/--
If `splitIndPred` is `true`, `grind` performs case-splitting on inductive predicates.
Otherwise, it performs case-splitting only on types marked with `[grind cases]` attribute. -/
splitIndPred : Bool := false
/-- By default, `grind` halts as soon as it encounters a sub-goal where no further progress can be made. -/
failures : Nat := 1
/-- Maximum number of heartbeats (in thousands) the canonicalizer can spend per definitional equality test. -/
canonHeartbeats : Nat := 1000
/-- If `ext` is `true`, `grind` uses extensionality theorems available in the environment. -/
ext : Bool := true
deriving Inhabited, BEq
end Lean.Grind
namespace Lean.Parser.Tactic
/-!
`grind` tactic and related tactics.
-/
syntax grindErase := "-" ident
syntax grindLemma := (Attr.grindMod)? ident
syntax grindParam := grindErase <|> grindLemma
syntax (name := grind)
"grind" optConfig (&" only")?
(" [" withoutPosition(grindParam,*) "]")?
("on_failure " term)? : tactic
syntax (name := grindTrace)
"grind?" optConfig (&" only")?
(" [" withoutPosition(grindParam,*) "]")?
("on_failure " term)? : tactic
end Lean.Parser.Tactic