-
Notifications
You must be signed in to change notification settings - Fork 433
/
Do.lean
1816 lines (1629 loc) · 76.9 KB
/
Do.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
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
492
493
494
495
496
497
498
499
500
501
502
503
504
505
506
507
508
509
510
511
512
513
514
515
516
517
518
519
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
560
561
562
563
564
565
566
567
568
569
570
571
572
573
574
575
576
577
578
579
580
581
582
583
584
585
586
587
588
589
590
591
592
593
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
609
610
611
612
613
614
615
616
617
618
619
620
621
622
623
624
625
626
627
628
629
630
631
632
633
634
635
636
637
638
639
640
641
642
643
644
645
646
647
648
649
650
651
652
653
654
655
656
657
658
659
660
661
662
663
664
665
666
667
668
669
670
671
672
673
674
675
676
677
678
679
680
681
682
683
684
685
686
687
688
689
690
691
692
693
694
695
696
697
698
699
700
701
702
703
704
705
706
707
708
709
710
711
712
713
714
715
716
717
718
719
720
721
722
723
724
725
726
727
728
729
730
731
732
733
734
735
736
737
738
739
740
741
742
743
744
745
746
747
748
749
750
751
752
753
754
755
756
757
758
759
760
761
762
763
764
765
766
767
768
769
770
771
772
773
774
775
776
777
778
779
780
781
782
783
784
785
786
787
788
789
790
791
792
793
794
795
796
797
798
799
800
801
802
803
804
805
806
807
808
809
810
811
812
813
814
815
816
817
818
819
820
821
822
823
824
825
826
827
828
829
830
831
832
833
834
835
836
837
838
839
840
841
842
843
844
845
846
847
848
849
850
851
852
853
854
855
856
857
858
859
860
861
862
863
864
865
866
867
868
869
870
871
872
873
874
875
876
877
878
879
880
881
882
883
884
885
886
887
888
889
890
891
892
893
894
895
896
897
898
899
900
901
902
903
904
905
906
907
908
909
910
911
912
913
914
915
916
917
918
919
920
921
922
923
924
925
926
927
928
929
930
931
932
933
934
935
936
937
938
939
940
941
942
943
944
945
946
947
948
949
950
951
952
953
954
955
956
957
958
959
960
961
962
963
964
965
966
967
968
969
970
971
972
973
974
975
976
977
978
979
980
981
982
983
984
985
986
987
988
989
990
991
992
993
994
995
996
997
998
999
1000
/-
Copyright (c) 2020 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura
-/
prelude
import Lean.Elab.Term
import Lean.Elab.BindersUtil
import Lean.Elab.PatternVar
import Lean.Elab.Quotation.Util
import Lean.Parser.Do
-- HACK: avoid code explosion until heuristics are improved
set_option compiler.reuse false
namespace Lean.Elab.Term
open Lean.Parser.Term
open Meta
open TSyntax.Compat
private def getDoSeqElems (doSeq : Syntax) : List Syntax :=
if doSeq.getKind == ``Parser.Term.doSeqBracketed then
doSeq[1].getArgs.toList.map fun arg => arg[0]
else if doSeq.getKind == ``Parser.Term.doSeqIndent then
doSeq[0].getArgs.toList.map fun arg => arg[0]
else
[]
private def getDoSeq (doStx : Syntax) : Syntax :=
doStx[1]
@[builtin_term_elab liftMethod] def elabLiftMethod : TermElab := fun stx _ =>
throwErrorAt stx "invalid use of `(<- ...)`, must be nested inside a 'do' expression"
/-- Return true if we should not lift `(<- ...)` actions nested in the syntax nodes with the given kind. -/
private def liftMethodDelimiter (k : SyntaxNodeKind) : Bool :=
k == ``Parser.Term.do ||
k == ``Parser.Term.doSeqIndent ||
k == ``Parser.Term.doSeqBracketed ||
k == ``Parser.Term.termReturn ||
k == ``Parser.Term.termUnless ||
k == ``Parser.Term.termTry ||
k == ``Parser.Term.termFor
/-- Given `stx` which is a `letPatDecl`, `letEqnsDecl`, or `letIdDecl`, return true if it has binders. -/
private def letDeclArgHasBinders (letDeclArg : Syntax) : Bool :=
let k := letDeclArg.getKind
if k == ``Parser.Term.letPatDecl then
false
else if k == ``Parser.Term.letEqnsDecl then
true
else if k == ``Parser.Term.letIdDecl then
-- letIdLhs := binderIdent >> checkWsBefore "expected space before binders" >> many (ppSpace >> letIdBinder)) >> optType
let binders := letDeclArg[1]
binders.getNumArgs > 0
else
false
/-- Return `true` if the given `letDecl` contains binders. -/
private def letDeclHasBinders (letDecl : Syntax) : Bool :=
letDeclArgHasBinders letDecl[0]
/-- Return true if we should generate an error message when lifting a method over this kind of syntax. -/
private def liftMethodForbiddenBinder (stx : Syntax) : Bool :=
let k := stx.getKind
-- TODO: make this extensible in the future.
if k == ``Parser.Term.fun || k == ``Parser.Term.matchAlts ||
k == ``Parser.Term.doLetRec || k == ``Parser.Term.letrec then
-- It is never ok to lift over this kind of binder
true
-- The following kinds of `let`-expressions require extra checks to decide whether they contain binders or not
else if k == ``Parser.Term.let then
letDeclHasBinders stx[1]
else if k == ``Parser.Term.doLet then
letDeclHasBinders stx[2]
else if k == ``Parser.Term.doLetArrow then
letDeclArgHasBinders stx[2]
else
false
-- TODO: we must track whether we are inside a quotation or not.
private partial def hasLiftMethod : Syntax → Bool
| Syntax.node _ k args =>
if liftMethodDelimiter k then false
-- NOTE: We don't check for lifts in quotations here, which doesn't break anything but merely makes this rare case a
-- bit slower
else if k == ``Parser.Term.liftMethod then true
-- For `pure` if-then-else, we only lift `(<- ...)` occurring in the condition.
else if k == ``termDepIfThenElse || k == ``termIfThenElse then args.size >= 2 && hasLiftMethod args[1]!
else args.any hasLiftMethod
| _ => false
structure ExtractMonadResult where
m : Expr
returnType : Expr
expectedType : Expr
private def mkUnknownMonadResult : MetaM ExtractMonadResult := do
let u ← mkFreshLevelMVar
let v ← mkFreshLevelMVar
let m ← mkFreshExprMVar (← mkArrow (mkSort (mkLevelSucc u)) (mkSort (mkLevelSucc v)))
let returnType ← mkFreshExprMVar (mkSort (mkLevelSucc u))
return { m, returnType, expectedType := mkApp m returnType }
private partial def extractBind (expectedType? : Option Expr) : TermElabM ExtractMonadResult := do
let some expectedType := expectedType? | mkUnknownMonadResult
let extractStep? (type : Expr) : MetaM (Option ExtractMonadResult) := do
let .app m returnType := type | return none
try
let bindInstType ← mkAppM ``Bind #[m]
discard <| Meta.synthInstance bindInstType
return some { m, returnType, expectedType }
catch _ =>
return none
let rec extract? (type : Expr) : MetaM (Option ExtractMonadResult) := do
match (← extractStep? type) with
| some r => return r
| none =>
let typeNew ← whnfCore type
if typeNew != type then
extract? typeNew
else
if typeNew.getAppFn.isMVar then
mkUnknownMonadResult
else match (← unfoldDefinition? typeNew) with
| some typeNew => extract? typeNew
| none => return none
match (← extract? expectedType) with
| some r => return r
| none => throwError "invalid `do` notation, expected type is not a monad application{indentExpr expectedType}\nYou can use the `do` notation in pure code by writing `Id.run do` instead of `do`, where `Id` is the identity monad."
namespace Do
abbrev Var := Syntax -- TODO: should be `Ident`
/-- A `doMatch` alternative. `vars` is the array of variables declared by `patterns`. -/
structure Alt (σ : Type) where
ref : Syntax
vars : Array Var
patterns : Syntax
rhs : σ
deriving Inhabited
/-- A `doMatchExpr` alternative. -/
structure AltExpr (σ : Type) where
ref : Syntax
var? : Option Var
funName : Syntax
pvars : Array Syntax
rhs : σ
deriving Inhabited
def AltExpr.vars (alt : AltExpr σ) : Array Var := Id.run do
let mut vars := #[]
if let some var := alt.var? then
vars := vars.push var
for pvar in alt.pvars do
match pvar with
| `(_) => pure ()
| _ => vars := vars.push pvar
return vars
/--
Auxiliary datastructure for representing a `do` code block, and compiling "reassignments" (e.g., `x := x + 1`).
We convert `Code` into a `Syntax` term representing the:
- `do`-block, or
- the visitor argument for the `forIn` combinator.
We say the following constructors are terminals:
- `break`: for interrupting a `for x in s`
- `continue`: for interrupting the current iteration of a `for x in s`
- `return e`: for returning `e` as the result for the whole `do` computation block
- `action a`: for executing action `a` as a terminal
- `ite`: if-then-else
- `match`: pattern matching
- `jmp` a goto to a join-point
We say the terminals `break`, `continue`, `action`, and `return` are "exit points"
Note that, `return e` is not equivalent to `action (pure e)`. Here is an example:
```
def f (x : Nat) : IO Unit := do
if x == 0 then
return ()
IO.println "hello"
```
Executing `#eval f 0` will not print "hello". Now, consider
```
def g (x : Nat) : IO Unit := do
if x == 0 then
pure ()
IO.println "hello"
```
The `if` statement is essentially a noop, and "hello" is printed when we execute `g 0`.
- `decl` represents all declaration-like `doElem`s (e.g., `let`, `have`, `let rec`).
The field `stx` is the actual `doElem`,
`vars` is the array of variables declared by it, and `cont` is the next instruction in the `do` code block.
`vars` is an array since we have declarations such as `let (a, b) := s`.
- `reassign` is an reassignment-like `doElem` (e.g., `x := x + 1`).
- `joinpoint` is a join point declaration: an auxiliary `let`-declaration used to represent the control-flow.
- `seq a k` executes action `a`, ignores its result, and then executes `k`.
We also store the do-elements `dbg_trace` and `assert!` as actions in a `seq`.
A code block `C` is well-formed if
- For every `jmp ref j as` in `C`, there is a `joinpoint j ps b k` and `jmp ref j as` is in `k`, and
`ps.size == as.size` -/
inductive Code where
| decl (xs : Array Var) (doElem : Syntax) (k : Code)
| reassign (xs : Array Var) (doElem : Syntax) (k : Code)
/-- The Boolean value in `params` indicates whether we should use `(x : typeof! x)` when generating term Syntax or not -/
| joinpoint (name : Name) (params : Array (Var × Bool)) (body : Code) (k : Code)
| seq (action : Syntax) (k : Code)
| action (action : Syntax)
| break (ref : Syntax)
| continue (ref : Syntax)
| return (ref : Syntax) (val : Syntax)
/-- Recall that an if-then-else may declare a variable using `optIdent` for the branches `thenBranch` and `elseBranch`. We store the variable name at `var?`. -/
| ite (ref : Syntax) (h? : Option Var) (optIdent : Syntax) (cond : Syntax) (thenBranch : Code) (elseBranch : Code)
| match (ref : Syntax) (gen : Syntax) (discrs : Syntax) (optMotive : Syntax) (alts : Array (Alt Code))
| matchExpr (ref : Syntax) (meta : Bool) (discr : Syntax) (alts : Array (AltExpr Code)) (elseBranch : Code)
| jmp (ref : Syntax) (jpName : Name) (args : Array Syntax)
deriving Inhabited
def Code.getRef? : Code → Option Syntax
| .decl _ doElem _ => doElem
| .reassign _ doElem _ => doElem
| .joinpoint .. => none
| .seq a _ => a
| .action a => a
| .break ref => ref
| .continue ref => ref
| .return ref _ => ref
| .ite ref .. => ref
| .match ref .. => ref
| .matchExpr ref .. => ref
| .jmp ref .. => ref
abbrev VarSet := RBMap Name Syntax Name.cmp
/-- A code block, and the collection of variables updated by it. -/
structure CodeBlock where
code : Code
uvars : VarSet := {} -- set of variables updated by `code`
private def varSetToArray (s : VarSet) : Array Var :=
s.fold (fun xs _ x => xs.push x) #[]
private def varsToMessageData (vars : Array Var) : MessageData :=
MessageData.joinSep (vars.toList.map fun n => MessageData.ofName (n.getId.simpMacroScopes)) " "
partial def CodeBlocl.toMessageData (codeBlock : CodeBlock) : MessageData :=
let us := MessageData.ofList <| (varSetToArray codeBlock.uvars).toList.map MessageData.ofSyntax
let rec loop : Code → MessageData
| .decl xs _ k => m!"let {varsToMessageData xs} := ...\n{loop k}"
| .reassign xs _ k => m!"{varsToMessageData xs} := ...\n{loop k}"
| .joinpoint n ps body k => m!"let {n.simpMacroScopes} {varsToMessageData (ps.map Prod.fst)} := {indentD (loop body)}\n{loop k}"
| .seq e k => m!"{e}\n{loop k}"
| .action e => e
| .ite _ _ _ c t e => m!"if {c} then {indentD (loop t)}\nelse{loop e}"
| .jmp _ j xs => m!"jmp {j.simpMacroScopes} {xs.toList}"
| .break _ => m!"break {us}"
| .continue _ => m!"continue {us}"
| .return _ v => m!"return {v} {us}"
| .match _ _ ds _ alts =>
m!"match {ds} with"
++ alts.foldl (init := m!"") fun acc alt => acc ++ m!"\n| {alt.patterns} => {loop alt.rhs}"
| .matchExpr _ meta d alts elseCode =>
let r := m!"match_expr {if meta then "" else "(meta := false)"} {d} with"
let r := r ++ alts.foldl (init := m!"") fun acc alt =>
let acc := acc ++ m!"\n| {if let some var := alt.var? then m!"{var}@" else ""}"
let acc := acc ++ m!"{alt.funName}"
let acc := acc ++ alt.pvars.foldl (init := m!"") fun acc pvar => acc ++ m!" {pvar}"
acc ++ m!" => {loop alt.rhs}"
r ++ m!"| _ => {loop elseCode}"
loop codeBlock.code
/-- Return true if the give code contains an exit point that satisfies `p` -/
partial def hasExitPointPred (c : Code) (p : Code → Bool) : Bool :=
let rec loop : Code → Bool
| .decl _ _ k => loop k
| .reassign _ _ k => loop k
| .joinpoint _ _ b k => loop b || loop k
| .seq _ k => loop k
| .ite _ _ _ _ t e => loop t || loop e
| .match _ _ _ _ alts => alts.any (loop ·.rhs)
| .matchExpr _ _ _ alts e => alts.any (loop ·.rhs) || loop e
| .jmp .. => false
| c => p c
loop c
def hasExitPoint (c : Code) : Bool :=
hasExitPointPred c fun _ => true
def hasReturn (c : Code) : Bool :=
hasExitPointPred c fun
| .return .. => true
| _ => false
def hasTerminalAction (c : Code) : Bool :=
hasExitPointPred c fun
| .action _ => true
| _ => false
def hasBreakContinue (c : Code) : Bool :=
hasExitPointPred c fun
| .break _ => true
| .continue _ => true
| _ => false
def hasBreakContinueReturn (c : Code) : Bool :=
hasExitPointPred c fun
| .break _ => true
| .continue _ => true
| .return _ _ => true
| _ => false
def mkAuxDeclFor {m} [Monad m] [MonadQuotation m] (e : Syntax) (mkCont : Syntax → m Code) : m Code := withRef e <| withFreshMacroScope do
let y ← `(y)
let doElem ← `(doElem| let y ← $e:term)
-- Add elaboration hint for producing sane error message
let y ← `(ensure_expected_type% "type mismatch, result value" $y)
let k ← mkCont y
return .decl #[y] doElem k
/-- Convert `action _ e` instructions in `c` into `let y ← e; jmp _ jp (xs y)`. -/
partial def convertTerminalActionIntoJmp (code : Code) (jp : Name) (xs : Array Var) : MacroM Code :=
let rec loop : Code → MacroM Code
| .decl xs stx k => return .decl xs stx (← loop k)
| .reassign xs stx k => return .reassign xs stx (← loop k)
| .joinpoint n ps b k => return .joinpoint n ps (← loop b) (← loop k)
| .seq e k => return .seq e (← loop k)
| .ite ref x? h c t e => return .ite ref x? h c (← loop t) (← loop e)
| .action e => mkAuxDeclFor e fun y =>
let ref := e
-- We jump to `jp` with xs **and** y
let jmpArgs := xs.push y
return Code.jmp ref jp jmpArgs
| .match ref g ds t alts =>
return .match ref g ds t (← alts.mapM fun alt => do pure { alt with rhs := (← loop alt.rhs) })
| .matchExpr ref meta d alts e => do
let alts ← alts.mapM fun alt => do pure { alt with rhs := (← loop alt.rhs) }
let e ← loop e
return .matchExpr ref meta d alts e
| c => return c
loop code
structure JPDecl where
name : Name
params : Array (Var × Bool)
body : Code
def attachJP (jpDecl : JPDecl) (k : Code) : Code :=
Code.joinpoint jpDecl.name jpDecl.params jpDecl.body k
def attachJPs (jpDecls : Array JPDecl) (k : Code) : Code :=
jpDecls.foldr attachJP k
def mkFreshJP (ps : Array (Var × Bool)) (body : Code) : TermElabM JPDecl := do
let ps ← if ps.isEmpty then
let y ← `(y)
pure #[(y.raw, false)]
else
pure ps
-- Remark: the compiler frontend implemented in C++ currently detects jointpoints created by
-- the "do" notation by testing the name. See hack at method `visit_let` at `lcnf.cpp`
-- We will remove this hack when we re-implement the compiler frontend in Lean.
let name ← mkFreshUserName `__do_jp
pure { name := name, params := ps, body := body }
def addFreshJP (ps : Array (Var × Bool)) (body : Code) : StateRefT (Array JPDecl) TermElabM Name := do
let jp ← mkFreshJP ps body
modify fun (jps : Array JPDecl) => jps.push jp
pure jp.name
def insertVars (rs : VarSet) (xs : Array Var) : VarSet :=
xs.foldl (fun rs x => rs.insert x.getId x) rs
def eraseVars (rs : VarSet) (xs : Array Var) : VarSet :=
xs.foldl (·.erase ·.getId) rs
def eraseOptVar (rs : VarSet) (x? : Option Var) : VarSet :=
match x? with
| none => rs
| some x => rs.insert x.getId x
/-- Create a new jointpoint for `c`, and jump to it with the variables `rs` -/
def mkSimpleJmp (ref : Syntax) (rs : VarSet) (c : Code) : StateRefT (Array JPDecl) TermElabM Code := do
let xs := varSetToArray rs
let jp ← addFreshJP (xs.map fun x => (x, true)) c
if xs.isEmpty then
let unit ← ``(Unit.unit)
return Code.jmp ref jp #[unit]
else
return Code.jmp ref jp xs
/-- Create a new joinpoint that takes `rs` and `val` as arguments. `val` must be syntax representing a pure value.
The body of the joinpoint is created using `mkJPBody yFresh`, where `yFresh`
is a fresh variable created by this method. -/
def mkJmp (ref : Syntax) (rs : VarSet) (val : Syntax) (mkJPBody : Syntax → MacroM Code) : StateRefT (Array JPDecl) TermElabM Code := do
let xs := varSetToArray rs
let args := xs.push val
let yFresh ← withRef ref `(y)
let ps := xs.map fun x => (x, true)
let ps := ps.push (yFresh, false)
let jpBody ← liftMacroM <| mkJPBody yFresh
let jp ← addFreshJP ps jpBody
return Code.jmp ref jp args
/-- `pullExitPointsAux rs c` auxiliary method for `pullExitPoints`, `rs` is the set of update variable in the current path. -/
partial def pullExitPointsAux (rs : VarSet) (c : Code) : StateRefT (Array JPDecl) TermElabM Code := do
match c with
| .decl xs stx k => return .decl xs stx (← pullExitPointsAux (eraseVars rs xs) k)
| .reassign xs stx k => return .reassign xs stx (← pullExitPointsAux (insertVars rs xs) k)
| .joinpoint j ps b k => return .joinpoint j ps (← pullExitPointsAux rs b) (← pullExitPointsAux rs k)
| .seq e k => return .seq e (← pullExitPointsAux rs k)
| .ite ref x? o c t e => return .ite ref x? o c (← pullExitPointsAux (eraseOptVar rs x?) t) (← pullExitPointsAux (eraseOptVar rs x?) e)
| .jmp .. => return c
| .break ref => mkSimpleJmp ref rs (.break ref)
| .continue ref => mkSimpleJmp ref rs (.continue ref)
| .return ref val => mkJmp ref rs val (fun y => return .return ref y)
| .action e =>
-- We use `mkAuxDeclFor` because `e` is not pure.
mkAuxDeclFor e fun y =>
let ref := e
mkJmp ref rs y (fun yFresh => return .action (← ``(Pure.pure $yFresh)))
| .match ref g ds t alts =>
let alts ← alts.mapM fun alt => do pure { alt with rhs := (← pullExitPointsAux (eraseVars rs alt.vars) alt.rhs) }
return .match ref g ds t alts
| .matchExpr ref meta d alts e =>
let alts ← alts.mapM fun alt => do pure { alt with rhs := (← pullExitPointsAux (eraseVars rs alt.vars) alt.rhs) }
let e ← pullExitPointsAux rs e
return .matchExpr ref meta d alts e
/--
Auxiliary operation for adding new variables to the collection of updated variables in a CodeBlock.
When a new variable is not already in the collection, but is shadowed by some declaration in `c`,
we create auxiliary join points to make sure we preserve the semantics of the code block.
Example: suppose we have the code block `print x; let x := 10; return x`. And we want to extend it
with the reassignment `x := x + 1`. We first use `pullExitPoints` to create
```
let jp (x!1) := return x!1;
print x;
let x := 10;
jmp jp x
```
and then we add the reassignment
```
x := x + 1
let jp (x!1) := return x!1;
print x;
let x := 10;
jmp jp x
```
Note that we created a fresh variable `x!1` to avoid accidental name capture.
As another example, consider
```
print x;
let x := 10
y := y + 1;
return x;
```
We transform it into
```
let jp (y x!1) := return x!1;
print x;
let x := 10
y := y + 1;
jmp jp y x
```
and then we add the reassignment as in the previous example.
We need to include `y` in the jump, because each exit point is implicitly returning the set of
update variables.
We implement the method as follows. Let `us` be `c.uvars`, then
1- for each `return _ y` in `c`, we create a join point
`let j (us y!1) := return y!1`
and replace the `return _ y` with `jmp us y`
2- for each `break`, we create a join point
`let j (us) := break`
and replace the `break` with `jmp us`.
3- Same as 2 for `continue`.
-/
def pullExitPoints (c : Code) : TermElabM Code := do
if hasExitPoint c then
let (c, jpDecls) ← (pullExitPointsAux {} c).run #[]
return attachJPs jpDecls c
else
return c
partial def extendUpdatedVarsAux (c : Code) (ws : VarSet) : TermElabM Code :=
let rec update (c : Code) : TermElabM Code := do
match c with
| .joinpoint j ps b k => return .joinpoint j ps (← update b) (← update k)
| .seq e k => return .seq e (← update k)
| .match ref g ds t alts =>
if alts.any fun alt => alt.vars.any fun x => ws.contains x.getId then
-- If a pattern variable is shadowing a variable in ws, we `pullExitPoints`
pullExitPoints c
else
return .match ref g ds t (← alts.mapM fun alt => do pure { alt with rhs := (← update alt.rhs) })
| .matchExpr ref meta d alts e =>
if alts.any fun alt => alt.vars.any fun x => ws.contains x.getId then
-- If a pattern variable is shadowing a variable in ws, we `pullExitPoints`
pullExitPoints c
else
let alts ← alts.mapM fun alt => do pure { alt with rhs := (← update alt.rhs) }
let e ← update e
return .matchExpr ref meta d alts e
| .ite ref none o c t e => return .ite ref none o c (← update t) (← update e)
| .ite ref (some h) o cond t e =>
if ws.contains h.getId then
-- if the `h` at `if h : c then t else e` shadows a variable in `ws`, we `pullExitPoints`
pullExitPoints c
else
return Code.ite ref (some h) o cond (← update t) (← update e)
| .reassign xs stx k => return .reassign xs stx (← update k)
| .decl xs stx k => do
if xs.any fun x => ws.contains x.getId then
-- One the declared variables is shadowing a variable in `ws`
pullExitPoints c
else
return .decl xs stx (← update k)
| c => return c
update c
/--
Extend the set of updated variables. It assumes `ws` is a super set of `c.uvars`.
We **cannot** simply update the field `c.uvars`, because `c` may have shadowed some variable in `ws`.
See discussion at `pullExitPoints`.
-/
partial def extendUpdatedVars (c : CodeBlock) (ws : VarSet) : TermElabM CodeBlock := do
if ws.any fun x _ => !c.uvars.contains x then
-- `ws` contains a variable that is not in `c.uvars`, but in `c.dvars` (i.e., it has been shadowed)
pure { code := (← extendUpdatedVarsAux c.code ws), uvars := ws }
else
pure { c with uvars := ws }
private def union (s₁ s₂ : VarSet) : VarSet :=
s₁.fold (·.insert ·) s₂
/--
Given two code blocks `c₁` and `c₂`, make sure they have the same set of updated variables.
Let `ws` the union of the updated variables in `c₁‵ and ‵c₂`.
We use `extendUpdatedVars c₁ ws` and `extendUpdatedVars c₂ ws`
-/
def homogenize (c₁ c₂ : CodeBlock) : TermElabM (CodeBlock × CodeBlock) := do
let ws := union c₁.uvars c₂.uvars
let c₁ ← extendUpdatedVars c₁ ws
let c₂ ← extendUpdatedVars c₂ ws
pure (c₁, c₂)
/--
Extending code blocks with variable declarations: `let x : t := v` and `let x : t ← v`.
We remove `x` from the collection of updated variables.
Remark: `stx` is the syntax for the declaration (e.g., `letDecl`), and `xs` are the variables
declared by it. It is an array because we have let-declarations that declare multiple variables.
Example: `let (x, y) := t`
-/
def mkVarDeclCore (xs : Array Var) (stx : Syntax) (c : CodeBlock) : CodeBlock := {
code := Code.decl xs stx c.code,
uvars := eraseVars c.uvars xs
}
/--
Extending code blocks with reassignments: `x : t := v` and `x : t ← v`.
Remark: `stx` is the syntax for the declaration (e.g., `letDecl`), and `xs` are the variables
declared by it. It is an array because we have let-declarations that declare multiple variables.
Example: `(x, y) ← t`
-/
def mkReassignCore (xs : Array Var) (stx : Syntax) (c : CodeBlock) : TermElabM CodeBlock := do
let us := c.uvars
let ws := insertVars us xs
-- If `xs` contains a new updated variable, then we must use `extendUpdatedVars`.
-- See discussion at `pullExitPoints`
let code ← if xs.any fun x => !us.contains x.getId then extendUpdatedVarsAux c.code ws else pure c.code
pure { code := .reassign xs stx code, uvars := ws }
def mkSeq (action : Syntax) (c : CodeBlock) : CodeBlock :=
{ c with code := .seq action c.code }
def mkTerminalAction (action : Syntax) : CodeBlock :=
{ code := .action action }
def mkReturn (ref : Syntax) (val : Syntax) : CodeBlock :=
{ code := .return ref val }
def mkBreak (ref : Syntax) : CodeBlock :=
{ code := .break ref }
def mkContinue (ref : Syntax) : CodeBlock :=
{ code := .continue ref }
def mkIte (ref : Syntax) (optIdent : Syntax) (cond : Syntax) (thenBranch : CodeBlock) (elseBranch : CodeBlock) : TermElabM CodeBlock := do
let x? := optIdent.getOptional?
let (thenBranch, elseBranch) ← homogenize thenBranch elseBranch
return {
code := .ite ref x? optIdent cond thenBranch.code elseBranch.code,
uvars := thenBranch.uvars,
}
private def mkUnit : MacroM Syntax :=
``((⟨⟩ : PUnit))
private def mkPureUnit : MacroM Syntax :=
``(pure PUnit.unit)
def mkPureUnitAction : MacroM CodeBlock := do
return mkTerminalAction (← mkPureUnit)
def mkUnless (cond : Syntax) (c : CodeBlock) : MacroM CodeBlock := do
let thenBranch ← mkPureUnitAction
return { c with code := .ite (← getRef) none mkNullNode cond thenBranch.code c.code }
def mkMatch (ref : Syntax) (genParam : Syntax) (discrs : Syntax) (optMotive : Syntax) (alts : Array (Alt CodeBlock)) : TermElabM CodeBlock := do
-- nary version of homogenize
let ws := alts.foldl (union · ·.rhs.uvars) {}
let alts ← alts.mapM fun alt => do
let rhs ← extendUpdatedVars alt.rhs ws
return { ref := alt.ref, vars := alt.vars, patterns := alt.patterns, rhs := rhs.code : Alt Code }
return { code := .match ref genParam discrs optMotive alts, uvars := ws }
def mkMatchExpr (ref : Syntax) (meta : Bool) (discr : Syntax) (alts : Array (AltExpr CodeBlock)) (elseBranch : CodeBlock) : TermElabM CodeBlock := do
-- nary version of homogenize
let ws := alts.foldl (union · ·.rhs.uvars) {}
let ws := union ws elseBranch.uvars
let alts ← alts.mapM fun alt => do
let rhs ← extendUpdatedVars alt.rhs ws
return { alt with rhs := rhs.code : AltExpr Code }
let elseBranch ← extendUpdatedVars elseBranch ws
return { code := .matchExpr ref meta discr alts elseBranch.code, uvars := ws }
/-- Return a code block that executes `terminal` and then `k` with the value produced by `terminal`.
This method assumes `terminal` is a terminal -/
def concat (terminal : CodeBlock) (kRef : Syntax) (y? : Option Var) (k : CodeBlock) : TermElabM CodeBlock := do
unless hasTerminalAction terminal.code do
throwErrorAt kRef "`do` element is unreachable"
let (terminal, k) ← homogenize terminal k
let xs := varSetToArray k.uvars
let y ← match y? with | some y => pure y | none => `(y)
let ps := xs.map fun x => (x, true)
let ps := ps.push (y, false)
let jpDecl ← mkFreshJP ps k.code
let jp := jpDecl.name
let terminal ← liftMacroM <| convertTerminalActionIntoJmp terminal.code jp xs
return { code := attachJP jpDecl terminal, uvars := k.uvars }
def getLetIdDeclVars (letIdDecl : Syntax) : Array Var :=
if letIdDecl[0].isIdent then
#[letIdDecl[0]]
else
#[]
-- support both regular and syntax match
def getPatternVarsEx (pattern : Syntax) : TermElabM (Array Var) :=
getPatternVars pattern <|>
Quotation.getPatternVars pattern
def getPatternsVarsEx (patterns : Array Syntax) : TermElabM (Array Var) :=
getPatternsVars patterns <|>
Quotation.getPatternsVars patterns
def getLetPatDeclVars (letPatDecl : Syntax) : TermElabM (Array Var) := do
let pattern := letPatDecl[0]
getPatternVarsEx pattern
def getLetEqnsDeclVars (letEqnsDecl : Syntax) : Array Var :=
if letEqnsDecl[0].isIdent then
#[letEqnsDecl[0]]
else
#[]
def getLetDeclVars (letDecl : Syntax) : TermElabM (Array Var) := do
let arg := letDecl[0]
if arg.getKind == ``Parser.Term.letIdDecl then
return getLetIdDeclVars arg
else if arg.getKind == ``Parser.Term.letPatDecl then
getLetPatDeclVars arg
else if arg.getKind == ``Parser.Term.letEqnsDecl then
return getLetEqnsDeclVars arg
else
throwError "unexpected kind of let declaration"
def getDoLetVars (doLet : Syntax) : TermElabM (Array Var) :=
-- leading_parser "let " >> optional "mut " >> letDecl
getLetDeclVars doLet[2]
def getDoHaveVars : Syntax → TermElabM (Array Var)
-- NOTE: `hygieneInfo` case should come first as `id` will match anything else
| `(doElem| have $info:hygieneInfo $_params* $[$_:typeSpec]? := $_val)
| `(doElem| have $info:hygieneInfo $_params* $[$_:typeSpec]? $_eqns:matchAlts) =>
return #[HygieneInfo.mkIdent info `this]
| `(doElem| have $id $_params* $[$_:typeSpec]? := $_val)
| `(doElem| have $id $_params* $[$_:typeSpec]? $_eqns:matchAlts) => return #[id]
| `(doElem| have $pat:letPatDecl) => getLetPatDeclVars pat
| _ => throwError "unexpected kind of have declaration"
def getDoLetRecVars (doLetRec : Syntax) : TermElabM (Array Var) := do
-- letRecDecls is an array of `(group (optional attributes >> letDecl))`
let letRecDecls := doLetRec[1][0].getSepArgs
let letDecls := letRecDecls.map fun p => p[2]
let mut allVars := #[]
for letDecl in letDecls do
let vars ← getLetDeclVars letDecl
allVars := allVars ++ vars
return allVars
-- ident >> optType >> leftArrow >> termParser
def getDoIdDeclVar (doIdDecl : Syntax) : Var :=
doIdDecl[0]
-- termParser >> leftArrow >> termParser >> optional (" | " >> termParser)
def getDoPatDeclVars (doPatDecl : Syntax) : TermElabM (Array Var) := do
let pattern := doPatDecl[0]
getPatternVarsEx pattern
-- leading_parser "let " >> optional "mut " >> (doIdDecl <|> doPatDecl)
def getDoLetArrowVars (doLetArrow : Syntax) : TermElabM (Array Var) := do
let decl := doLetArrow[2]
if decl.getKind == ``Parser.Term.doIdDecl then
return #[getDoIdDeclVar decl]
else if decl.getKind == ``Parser.Term.doPatDecl then
getDoPatDeclVars decl
else
throwError "unexpected kind of `do` declaration"
def getDoReassignVars (doReassign : Syntax) : TermElabM (Array Var) := do
let arg := doReassign[0]
if arg.getKind == ``Parser.Term.letIdDecl then
return getLetIdDeclVars arg
else if arg.getKind == ``Parser.Term.letPatDecl then
getLetPatDeclVars arg
else
throwError "unexpected kind of reassignment"
def mkDoSeq (doElems : Array Syntax) : Syntax :=
mkNode `Lean.Parser.Term.doSeqIndent #[mkNullNode <| doElems.map fun doElem => mkNullNode #[doElem, mkNullNode]]
/--
If the given syntax is a `doIf`, return an equivalent `doIf` that has an `else` but no `else if`s or `if let`s. -/
private def expandDoIf? (stx : Syntax) : MacroM (Option Syntax) := match stx with
| `(doElem|if $_:doIfProp then $_ else $_) => pure none
| `(doElem|if $cond:doIfCond then $t $[else if $conds:doIfCond then $ts]* $[else $e?]?) => withRef stx do
let mut e := e?.getD (← `(doSeq|pure PUnit.unit))
let mut eIsSeq := true
for (cond, t) in Array.zip (conds.reverse.push cond) (ts.reverse.push t) do
e ← if eIsSeq then pure e else `(doSeq|$e:doElem)
e ← match cond with
| `(doIfCond|let $pat := $d) => `(doElem| match $d:term with | $pat:term => $t | _ => $e)
| `(doIfCond|let $pat ← $d) => `(doElem| match ← $d with | $pat:term => $t | _ => $e)
| `(doIfCond|$cond:doIfProp) => `(doElem| if $cond:doIfProp then $t else $e)
| _ => `(doElem| if $(Syntax.missing) then $t else $e)
eIsSeq := false
return some e
| _ => pure none
/--
If the given syntax is a `doLetExpr` or `doLetMetaExpr`, return an equivalent `doIf` that has an `else` but no `else if`s or `if let`s. -/
private def expandDoLetExpr? (stx : Syntax) (doElems : List Syntax) : MacroM (Option Syntax) := match stx with
| `(doElem| let_expr $pat:matchExprPat := $discr:term | $elseBranch:doSeq) =>
return some (← `(doElem| match_expr (meta := false) $discr:term with
| $pat:matchExprPat => $(mkDoSeq doElems.toArray)
| _ => $elseBranch))
| `(doElem| let_expr $pat:matchExprPat ← $discr:term | $elseBranch:doSeq) =>
return some (← `(doElem| match_expr $discr:term with
| $pat:matchExprPat => $(mkDoSeq doElems.toArray)
| _ => $elseBranch))
| _ => return none
structure DoIfView where
ref : Syntax
optIdent : Syntax
cond : Syntax
thenBranch : Syntax
elseBranch : Syntax
/-- This method assumes `expandDoIf?` is not applicable. -/
private def mkDoIfView (doIf : Syntax) : DoIfView := {
ref := doIf
optIdent := doIf[1][0]
cond := doIf[1][1]
thenBranch := doIf[3]
elseBranch := doIf[5][1]
}
/--
We use `MProd` instead of `Prod` to group values when expanding the
`do` notation. `MProd` is a universe monomorphic product.
The motivation is to generate simpler universe constraints in code
that was not written by the user.
Note that we are not restricting the macro power since the
`Bind.bind` combinator already forces values computed by monadic
actions to be in the same universe.
-/
private def mkTuple (elems : Array Syntax) : MacroM Syntax := do
if elems.size == 0 then
mkUnit
else if elems.size == 1 then
return elems[0]!
else
elems.extract 0 (elems.size - 1) |>.foldrM (init := elems.back!) fun elem tuple =>
``(MProd.mk $elem $tuple)
/-- Return `some action` if `doElem` is a `doExpr <action>`-/
def isDoExpr? (doElem : Syntax) : Option Syntax :=
if doElem.getKind == ``Parser.Term.doExpr then
some doElem[0]
else
none
/--
Given `uvars := #[a_1, ..., a_n, a_{n+1}]` construct term
```
let a_1 := x.1
let x := x.2
let a_2 := x.1
let x := x.2
...
let a_n := x.1
let a_{n+1} := x.2
body
```
Special cases
- `uvars := #[]` => `body`
- `uvars := #[a]` => `let a := x; body`
We use this method when expanding the `for-in` notation.
-/
private def destructTuple (uvars : Array Var) (x : Syntax) (body : Syntax) : MacroM Syntax := do
if uvars.size == 0 then
return body
else if uvars.size == 1 then
`(let $(uvars[0]!):ident := $x; $body)
else
destruct uvars.toList x body
where
destruct (as : List Var) (x : Syntax) (body : Syntax) : MacroM Syntax := do
match as with
| [a, b] => `(let $a:ident := $x.1; let $b:ident := $x.2; $body)
| a :: as => withFreshMacroScope do
let rest ← destruct as (← `(x)) body
`(let $a:ident := $x.1; let x := $x.2; $rest)
| _ => unreachable!
/-!
The procedure `ToTerm.run` converts a `CodeBlock` into a `Syntax` term.
We use this method to convert
1- The `CodeBlock` for a root `do ...` term into a `Syntax` term. This kind of
`CodeBlock` never contains `break` nor `continue`. Moreover, the collection
of updated variables is not packed into the result.
Thus, we have two kinds of exit points
- `Code.action e` which is converted into `e`
- `Code.return _ e` which is converted into `pure e`
We use `Kind.regular` for this case.
2- The `CodeBlock` for `b` at `for x in xs do b`. In this case, we need to generate
a `Syntax` term representing a function for the `xs.forIn` combinator.
a) If `b` contain a `Code.return _ a` exit point. The generated `Syntax` term
has type `m (ForInStep (Option α × σ))`, where `a : α`, and the `σ` is the type
of the tuple of variables reassigned by `b`.
We use `Kind.forInWithReturn` for this case
b) If `b` does not contain a `Code.return _ a` exit point. Then, the generated
`Syntax` term has type `m (ForInStep σ)`.
We use `Kind.forIn` for this case.
3- The `CodeBlock` `c` for a `do` sequence nested in a monadic combinator (e.g., `MonadExcept.tryCatch`).
The generated `Syntax` term for `c` must inform whether `c` "exited" using `Code.action`, `Code.return`,
`Code.break` or `Code.continue`. We use the auxiliary types `DoResult`s for storing this information.
For example, the auxiliary type `DoResultPBC α σ` is used for a code block that exits with `Code.action`,
**and** `Code.break`/`Code.continue`, `α` is the type of values produced by the exit `action`, and
`σ` is the type of the tuple of reassigned variables.
The type `DoResult α β σ` is usedf for code blocks that exit with
`Code.action`, `Code.return`, **and** `Code.break`/`Code.continue`, `β` is the type of the returned values.
We don't use `DoResult α β σ` for all cases because:
a) The elaborator would not be able to infer all type parameters without extra annotations. For example,
if the code block does not contain `Code.return _ _`, the elaborator will not be able to infer `β`.
b) We need to pattern match on the result produced by the combinator (e.g., `MonadExcept.tryCatch`),
but we don't want to consider "unreachable" cases.
We do not distinguish between cases that contain `break`, but not `continue`, and vice versa.
When listing all cases, we use `a` to indicate the code block contains `Code.action _`, `r` for `Code.return _ _`,
and `b/c` for a code block that contains `Code.break _` or `Code.continue _`.
- `a`: `Kind.regular`, type `m (α × σ)`
- `r`: `Kind.regular`, type `m (α × σ)`
Note that the code that pattern matches on the result will behave differently in this case.
It produces `return a` for this case, and `pure a` for the previous one.
- `b/c`: `Kind.nestedBC`, type `m (DoResultBC σ)`
- `a` and `r`: `Kind.nestedPR`, type `m (DoResultPR α β σ)`
- `a` and `bc`: `Kind.nestedSBC`, type `m (DoResultSBC α σ)`
- `r` and `bc`: `Kind.nestedSBC`, type `m (DoResultSBC α σ)`
Again the code that pattern matches on the result will behave differently in this case and
the previous one. It produces `return a` for the constructor `DoResultSPR.pureReturn a u` for
this case, and `pure a` for the previous case.
- `a`, `r`, `b/c`: `Kind.nestedPRBC`, type type `m (DoResultPRBC α β σ)`
Here is the recipe for adding new combinators with nested `do`s.
Example: suppose we want to support `repeat doSeq`. Assuming we have `repeat : m α → m α`
1- Convert `doSeq` into `codeBlock : CodeBlock`
2- Create term `term` using `mkNestedTerm code m uvars a r bc` where
`code` is `codeBlock.code`, `uvars` is an array containing `codeBlock.uvars`,
`m` is a `Syntax` representing the Monad, and
`a` is true if `code` contains `Code.action _`,
`r` is true if `code` contains `Code.return _ _`,
`bc` is true if `code` contains `Code.break _` or `Code.continue _`.
Remark: for combinators such as `repeat` that take a single `doSeq`, all
arguments, but `m`, are extracted from `codeBlock`.
3- Create the term `repeat $term`
4- and then, convert it into a `doSeq` using `matchNestedTermResult ref (repeat $term) uvsar a r bc`
-/
/--
Helper method for annotating `term` with the raw syntax `ref`.
We use this method to implement finer-grained term infos for `do`-blocks.
We use `withRef term` to make sure the synthetic position for the `with_annotate_term` is equal
to the one for `term`. This is important for producing error messages when there is a type mismatch.
Consider the following example:
```
opaque f : IO Nat
def g : IO String := do
f
```
There is at type mismatch at `f`, but it is detected when elaborating the expanded term
containing the `with_annotate_term .. f`. The current `getRef` when this `annotate` is invoked
is not necessarily `f`. Actually, it is the whole `do`-block. By using `withRef` we ensure
the synthetic position for the `with_annotate_term ..` is equal to `term`.
Recall that synthetic positions are used when generating error messages.
-/
def annotate [Monad m] [MonadRef m] [MonadQuotation m] (ref : Syntax) (term : Syntax) : m Syntax :=
withRef term <| `(with_annotate_term $ref $term)
namespace ToTerm
inductive Kind where
| regular
| forIn
| forInWithReturn
| nestedBC
| nestedPR
| nestedSBC
| nestedPRBC
instance : Inhabited Kind := ⟨Kind.regular⟩
def Kind.isRegular : Kind → Bool
| .regular => true
| _ => false
structure Context where
/-- Syntax to reference the monad associated with the do notation. -/
m : Syntax
/-- Syntax to reference the result of the monadic computation performed by the do notation. -/
returnType : Syntax
uvars : Array Var
kind : Kind
abbrev M := ReaderT Context MacroM
def mkUVarTuple : M Syntax := do
let ctx ← read
mkTuple ctx.uvars
def returnToTerm (val : Syntax) : M Syntax := do
let ctx ← read
let u ← mkUVarTuple
match ctx.kind with
| .regular => if ctx.uvars.isEmpty then ``(Pure.pure $val) else ``(Pure.pure (MProd.mk $val $u))
| .forIn => ``(Pure.pure (ForInStep.done $u))
| .forInWithReturn => ``(Pure.pure (ForInStep.done (MProd.mk (some $val) $u)))
| .nestedBC => unreachable!
| .nestedPR => ``(Pure.pure (DoResultPR.«return» $val $u))
| .nestedSBC => ``(Pure.pure (DoResultSBC.«pureReturn» $val $u))
| .nestedPRBC => ``(Pure.pure (DoResultPRBC.«return» $val $u))
def continueToTerm : M Syntax := do
let ctx ← read
let u ← mkUVarTuple
match ctx.kind with
| .regular => unreachable!