-
Notifications
You must be signed in to change notification settings - Fork 465
/
Copy pathPassManager.lean
183 lines (149 loc) · 6.06 KB
/
PassManager.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
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
/-
Copyright (c) 2022 Henrik Böving. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Henrik Böving
-/
import Lean.Attributes
import Lean.Environment
import Lean.Meta.Basic
import Lean.Compiler.LCNF.CompilerM
namespace Lean.Compiler.LCNF
def Phase.toNat : Phase → Nat
| .base => 0
| .mono => 1
| .impure => 2
instance : LT Phase where
lt l r := l.toNat < r.toNat
instance : LE Phase where
le l r := l.toNat ≤ r.toNat
instance {p1 p2 : Phase} : Decidable (p1 < p2) := Nat.decLt p1.toNat p2.toNat
instance {p1 p2 : Phase} : Decidable (p1 ≤ p2) := Nat.decLe p1.toNat p2.toNat
@[simp] theorem Phase.le_refl (p : Phase) : p ≤ p := by
cases p <;> decide
/--
A single compiler `Pass`, consisting of the actual pass function operating
on the `Decl`s as well as meta information.
-/
structure Pass where
/--
Which occurrence of the pass in the pipeline this is.
Some passes, like simp, can occur multiple times in the pipeline.
For most passes this value does not matter.
-/
occurrence : Nat := 0
/--
Which phase this `Pass` is supposed to run in
-/
phase : Phase
/--
Resulting phase.
-/
phaseOut : Phase := phase
phaseInv : phaseOut ≥ phase := by simp
/--
The name of the `Pass`
-/
name : Name
/--
The actual pass function, operating on the `Decl`s.
-/
run : Array Decl → CompilerM (Array Decl)
instance : Inhabited Pass where
default := { phase := .base, name := default, run := fun decls => return decls }
/--
Can be used to install, remove, replace etc. passes by tagging a declaration
of type `PassInstaller` with the `cpass` attribute.
-/
structure PassInstaller where
/--
When the installer is run this function will receive a list of all
current `Pass`es and return a new one, this can modify the list (and
the `Pass`es contained within) in any way it wants.
-/
install : Array Pass → CoreM (Array Pass)
deriving Inhabited
/--
The `PassManager` used to store all `Pass`es that will be run within
pipeline.
-/
structure PassManager where
passes : Array Pass
deriving Inhabited
instance : ToString Phase where
toString
| .base => "base"
| .mono => "mono"
| .impure => "impure"
namespace Pass
def mkPerDeclaration (name : Name) (run : Decl → CompilerM Decl) (phase : Phase) (occurrence : Nat := 0) : Pass where
occurrence := occurrence
phase := phase
name := name
run := fun xs => xs.mapM run
end Pass
namespace PassManager
def validate (manager : PassManager) : CoreM Unit := do
let mut current := .base
for pass in manager.passes do
if ¬(current ≤ pass.phase) then
throwError s!"{pass.name} has phase {pass.phase} but should at least have {current}"
current := pass.phase
def findHighestOccurrence (targetName : Name) (passes : Array Pass) : CoreM Nat := do
let mut highest := none
for pass in passes do
if pass.name == targetName then
highest := some pass.occurrence
let some val := highest | throwError s!"Could not find any occurrence of {targetName}"
return val
end PassManager
namespace PassInstaller
def installAtEnd (p : Pass) : PassInstaller where
install passes := return passes.push p
def append (passesNew : Array Pass) : PassInstaller where
install passes := return passes ++ passesNew
def withEachOccurrence (targetName : Name) (f : Nat → PassInstaller) : PassInstaller where
install passes := do
let highestOccurrence ← PassManager.findHighestOccurrence targetName passes
let mut passes := passes
for occurrence in [0:highestOccurrence+1] do
passes ← f occurrence |>.install passes
return passes
def installAfter (targetName : Name) (p : Pass → Pass) (occurrence : Nat := 0) : PassInstaller where
install passes :=
if let some idx := passes.findIdx? (fun p => p.name == targetName && p.occurrence == occurrence) then
let passUnderTest := passes[idx]!
return passes.insertAt! (idx + 1) (p passUnderTest)
else
throwError s!"Tried to insert pass after {targetName}, occurrence {occurrence} but {targetName} is not in the pass list"
def installAfterEach (targetName : Name) (p : Pass → Pass) : PassInstaller :=
withEachOccurrence targetName (installAfter targetName p ·)
def installBefore (targetName : Name) (p : Pass → Pass) (occurrence : Nat := 0): PassInstaller where
install passes :=
if let some idx := passes.findIdx? (fun p => p.name == targetName && p.occurrence == occurrence) then
let passUnderTest := passes[idx]!
return passes.insertAt! idx (p passUnderTest)
else
throwError s!"Tried to insert pass after {targetName}, occurrence {occurrence} but {targetName} is not in the pass list"
def installBeforeEachOccurrence (targetName : Name) (p : Pass → Pass) : PassInstaller :=
withEachOccurrence targetName (installBefore targetName p ·)
def replacePass (targetName : Name) (p : Pass → Pass) (occurrence : Nat := 0) : PassInstaller where
install passes := do
let some idx := passes.findIdx? (fun p => p.name == targetName && p.occurrence == occurrence) | throwError s!"Tried to replace {targetName}, occurrence {occurrence} but {targetName} is not in the pass list"
let target := passes[idx]!
let replacement := p target
return passes.set! idx replacement
def replaceEachOccurrence (targetName : Name) (p : Pass → Pass) : PassInstaller :=
withEachOccurrence targetName (replacePass targetName p ·)
def run (manager : PassManager) (installer : PassInstaller) : CoreM PassManager := do
return { manager with passes := (← installer.install manager.passes) }
private unsafe def getPassInstallerUnsafe (declName : Name) : CoreM PassInstaller := do
ofExcept <| (← getEnv).evalConstCheck PassInstaller (← getOptions) ``PassInstaller declName
@[implemented_by getPassInstallerUnsafe]
private opaque getPassInstaller (declName : Name) : CoreM PassInstaller
def runFromDecl (manager : PassManager) (declName : Name) : CoreM PassManager := do
let installer ← getPassInstaller declName
let newState ← installer.run manager
newState.validate
return newState
end PassInstaller
end Lean.Compiler.LCNF