forked from leanprover/lean4
-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathMeta.lean
1350 lines (1086 loc) · 46 KB
/
Meta.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) 2019 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Leonardo de Moura and Sebastian Ullrich
Additional goodies for writing macros
-/
prelude
import Init.Data.Array.Basic
import Init.Data.Option.BasicAux
namespace Lean
@[extern c inline "lean_box(LEAN_VERSION_MAJOR)"]
private opaque version.getMajor (u : Unit) : Nat
def version.major : Nat := version.getMajor ()
@[extern c inline "lean_box(LEAN_VERSION_MINOR)"]
private opaque version.getMinor (u : Unit) : Nat
def version.minor : Nat := version.getMinor ()
@[extern c inline "lean_box(LEAN_VERSION_PATCH)"]
private opaque version.getPatch (u : Unit) : Nat
def version.patch : Nat := version.getPatch ()
@[extern "lean_get_githash"]
opaque getGithash (u : Unit) : String
def githash : String := getGithash ()
@[extern c inline "LEAN_VERSION_IS_RELEASE"]
opaque version.getIsRelease (u : Unit) : Bool
def version.isRelease : Bool := version.getIsRelease ()
/-- Additional version description like "nightly-2018-03-11" -/
@[extern c inline "lean_mk_string(LEAN_SPECIAL_VERSION_DESC)"]
opaque version.getSpecialDesc (u : Unit) : String
def version.specialDesc : String := version.getSpecialDesc ()
def versionStringCore :=
toString version.major ++ "." ++ toString version.minor ++ "." ++ toString version.patch
def versionString :=
if version.specialDesc ≠ "" then
versionStringCore ++ "-" ++ version.specialDesc
else if version.isRelease then
versionStringCore
else
versionStringCore ++ ", commit " ++ githash
def origin :=
"leanprover/lean4"
def toolchain :=
if version.specialDesc ≠ "" then
if version.isRelease then
origin ++ ":" ++ versionStringCore ++ "-" ++ version.specialDesc
else
origin ++ ":" ++ version.specialDesc
else if version.isRelease then
origin ++ ":" ++ versionStringCore
else
""
@[extern c inline "LEAN_IS_STAGE0"]
opaque Internal.isStage0 (u : Unit) : Bool
/-- Valid identifier names -/
def isGreek (c : Char) : Bool :=
0x391 ≤ c.val && c.val ≤ 0x3dd
def isLetterLike (c : Char) : Bool :=
(0x3b1 ≤ c.val && c.val ≤ 0x3c9 && c.val ≠ 0x3bb) || -- Lower greek, but lambda
(0x391 ≤ c.val && c.val ≤ 0x3A9 && c.val ≠ 0x3A0 && c.val ≠ 0x3A3) || -- Upper greek, but Pi and Sigma
(0x3ca ≤ c.val && c.val ≤ 0x3fb) || -- Coptic letters
(0x1f00 ≤ c.val && c.val ≤ 0x1ffe) || -- Polytonic Greek Extended Character Set
(0x2100 ≤ c.val && c.val ≤ 0x214f) || -- Letter like block
(0x1d49c ≤ c.val && c.val ≤ 0x1d59f) -- Latin letters, Script, Double-struck, Fractur
def isNumericSubscript (c : Char) : Bool :=
0x2080 ≤ c.val && c.val ≤ 0x2089
def isSubScriptAlnum (c : Char) : Bool :=
isNumericSubscript c ||
(0x2090 ≤ c.val && c.val ≤ 0x209c) ||
(0x1d62 ≤ c.val && c.val ≤ 0x1d6a)
def isIdFirst (c : Char) : Bool :=
c.isAlpha || c = '_' || isLetterLike c
def isIdRest (c : Char) : Bool :=
c.isAlphanum || c = '_' || c = '\'' || c == '!' || c == '?' || isLetterLike c || isSubScriptAlnum c
def idBeginEscape := '«'
def idEndEscape := '»'
def isIdBeginEscape (c : Char) : Bool := c = idBeginEscape
def isIdEndEscape (c : Char) : Bool := c = idEndEscape
namespace Name
def getRoot : Name → Name
| anonymous => anonymous
| n@(str anonymous _) => n
| n@(num anonymous _) => n
| str n _ => getRoot n
| num n _ => getRoot n
@[export lean_is_inaccessible_user_name]
def isInaccessibleUserName : Name → Bool
| Name.str _ s => s.contains '✝' || s == "_inaccessible"
| Name.num p _ => isInaccessibleUserName p
| _ => false
def escapePart (s : String) : Option String :=
if s.length > 0 && isIdFirst (s.get 0) && (s.toSubstring.drop 1).all isIdRest then s
else if s.any isIdEndEscape then none
else some <| idBeginEscape.toString ++ s ++ idEndEscape.toString
-- NOTE: does not roundtrip even with `escape = true` if name is anonymous or contains numeric part or `idEndEscape`
variable (sep : String) (escape : Bool)
def toStringWithSep : Name → String
| anonymous => "[anonymous]"
| str anonymous s => maybeEscape s
| num anonymous v => toString v
| str n s => toStringWithSep n ++ sep ++ maybeEscape s
| num n v => toStringWithSep n ++ sep ++ Nat.repr v
where
maybeEscape s := if escape then escapePart s |>.getD s else s
protected def toString (n : Name) (escape := true) : String :=
-- never escape "prettified" inaccessible names or macro scopes or pseudo-syntax introduced by the delaborator
toStringWithSep "." (escape && !n.isInaccessibleUserName && !n.hasMacroScopes && !maybePseudoSyntax) n
where
maybePseudoSyntax :=
if let .str _ s := n.getRoot then
-- could be pseudo-syntax for loose bvar or universe mvar, output as is
"#".isPrefixOf s || "?".isPrefixOf s
else
false
instance : ToString Name where
toString n := n.toString
private def hasNum : Name → Bool
| anonymous => false
| num .. => true
| str p .. => hasNum p
protected def reprPrec (n : Name) (prec : Nat) : Std.Format :=
match n with
| anonymous => Std.Format.text "Lean.Name.anonymous"
| num p i => Repr.addAppParen ("Lean.Name.mkNum " ++ Name.reprPrec p max_prec ++ " " ++ repr i) prec
| str p s =>
if p.hasNum then
Repr.addAppParen ("Lean.Name.mkStr " ++ Name.reprPrec p max_prec ++ " " ++ repr s) prec
else
Std.Format.text "`" ++ n.toString
instance : Repr Name where
reprPrec := Name.reprPrec
def capitalize : Name → Name
| .str p s => .str p s.capitalize
| n => n
def replacePrefix : Name → Name → Name → Name
| anonymous, anonymous, newP => newP
| anonymous, _, _ => anonymous
| n@(str p s), queryP, newP => if n == queryP then newP else Name.mkStr (p.replacePrefix queryP newP) s
| n@(num p s), queryP, newP => if n == queryP then newP else Name.mkNum (p.replacePrefix queryP newP) s
/--
`eraseSuffix? n s` return `n'` if `n` is of the form `n == n' ++ s`.
-/
def eraseSuffix? : Name → Name → Option Name
| n, anonymous => some n
| str p s, str p' s' => if s == s' then eraseSuffix? p p' else none
| num p s, num p' s' => if s == s' then eraseSuffix? p p' else none
| _, _ => none
/-- Remove macros scopes, apply `f`, and put them back -/
@[inline] def modifyBase (n : Name) (f : Name → Name) : Name :=
if n.hasMacroScopes then
let view := extractMacroScopes n
{ view with name := f view.name }.review
else
f n
@[export lean_name_append_after]
def appendAfter (n : Name) (suffix : String) : Name :=
n.modifyBase fun
| str p s => Name.mkStr p (s ++ suffix)
| n => Name.mkStr n suffix
@[export lean_name_append_index_after]
def appendIndexAfter (n : Name) (idx : Nat) : Name :=
n.modifyBase fun
| str p s => Name.mkStr p (s ++ "_" ++ toString idx)
| n => Name.mkStr n ("_" ++ toString idx)
@[export lean_name_append_before]
def appendBefore (n : Name) (pre : String) : Name :=
n.modifyBase fun
| anonymous => Name.mkStr anonymous pre
| str p s => Name.mkStr p (pre ++ s)
| num p n => Name.mkNum (Name.mkStr p pre) n
protected theorem beq_iff_eq {m n : Name} : m == n ↔ m = n := by
show m.beq n ↔ _
induction m generalizing n <;> cases n <;> simp_all [Name.beq, And.comm]
instance : LawfulBEq Name where
eq_of_beq := Name.beq_iff_eq.1
rfl := Name.beq_iff_eq.2 rfl
instance : DecidableEq Name :=
fun a b => if h : a == b then .isTrue (by simp_all) else .isFalse (by simp_all)
end Name
structure NameGenerator where
namePrefix : Name := `_uniq
idx : Nat := 1
deriving Inhabited
namespace NameGenerator
@[inline] def curr (g : NameGenerator) : Name :=
Name.mkNum g.namePrefix g.idx
@[inline] def next (g : NameGenerator) : NameGenerator :=
{ g with idx := g.idx + 1 }
@[inline] def mkChild (g : NameGenerator) : NameGenerator × NameGenerator :=
({ namePrefix := Name.mkNum g.namePrefix g.idx, idx := 1 },
{ g with idx := g.idx + 1 })
end NameGenerator
class MonadNameGenerator (m : Type → Type) where
getNGen : m NameGenerator
setNGen : NameGenerator → m Unit
export MonadNameGenerator (getNGen setNGen)
def mkFreshId {m : Type → Type} [Monad m] [MonadNameGenerator m] : m Name := do
let ngen ← getNGen
let r := ngen.curr
setNGen ngen.next
pure r
instance monadNameGeneratorLift (m n : Type → Type) [MonadLift m n] [MonadNameGenerator m] : MonadNameGenerator n := {
getNGen := liftM (getNGen : m _),
setNGen := fun ngen => liftM (setNGen ngen : m _)
}
namespace Syntax
deriving instance Repr for Syntax.Preresolved
deriving instance Repr for Syntax
deriving instance Repr for TSyntax
abbrev Term := TSyntax `term
abbrev Command := TSyntax `command
protected abbrev Level := TSyntax `level
protected abbrev Tactic := TSyntax `tactic
abbrev Prec := TSyntax `prec
abbrev Prio := TSyntax `prio
abbrev Ident := TSyntax identKind
abbrev StrLit := TSyntax strLitKind
abbrev CharLit := TSyntax charLitKind
abbrev NameLit := TSyntax nameLitKind
abbrev ScientificLit := TSyntax scientificLitKind
abbrev NumLit := TSyntax numLitKind
abbrev HygieneInfo := TSyntax hygieneInfoKind
end Syntax
export Syntax (Term Command Prec Prio Ident StrLit CharLit NameLit ScientificLit NumLit HygieneInfo)
namespace TSyntax
instance : Coe (TSyntax [k]) (TSyntax (k :: ks)) where
coe stx := ⟨stx⟩
instance : Coe (TSyntax ks) (TSyntax (k' :: ks)) where
coe stx := ⟨stx⟩
instance : Coe Ident Term where
coe s := ⟨s.raw⟩
instance : CoeDep Term ⟨Syntax.ident info ss n res⟩ Ident where
coe := ⟨Syntax.ident info ss n res⟩
instance : Coe StrLit Term where
coe s := ⟨s.raw⟩
instance : Coe NameLit Term where
coe s := ⟨s.raw⟩
instance : Coe ScientificLit Term where
coe s := ⟨s.raw⟩
instance : Coe NumLit Term where
coe s := ⟨s.raw⟩
instance : Coe CharLit Term where
coe s := ⟨s.raw⟩
instance : Coe Ident Syntax.Level where
coe s := ⟨s.raw⟩
instance : Coe NumLit Prio where
coe s := ⟨s.raw⟩
instance : Coe NumLit Prec where
coe s := ⟨s.raw⟩
namespace Compat
scoped instance : CoeTail Syntax (TSyntax k) where
coe s := ⟨s⟩
scoped instance : CoeTail (Array Syntax) (TSyntaxArray k) where
coe := .mk
end Compat
end TSyntax
namespace Syntax
deriving instance BEq for Syntax.Preresolved
/-- Compare syntax structures modulo source info. -/
partial def structEq : Syntax → Syntax → Bool
| Syntax.missing, Syntax.missing => true
| Syntax.node _ k args, Syntax.node _ k' args' => k == k' && args.isEqv args' structEq
| Syntax.atom _ val, Syntax.atom _ val' => val == val'
| Syntax.ident _ rawVal val preresolved, Syntax.ident _ rawVal' val' preresolved' => rawVal == rawVal' && val == val' && preresolved == preresolved'
| _, _ => false
instance : BEq Lean.Syntax := ⟨structEq⟩
instance : BEq (Lean.TSyntax k) := ⟨(·.raw == ·.raw)⟩
partial def getTailInfo? : Syntax → Option SourceInfo
| atom info _ => info
| ident info .. => info
| node SourceInfo.none _ args =>
args.findSomeRev? getTailInfo?
| node info _ _ => info
| _ => none
def getTailInfo (stx : Syntax) : SourceInfo :=
stx.getTailInfo?.getD SourceInfo.none
def getTrailingSize (stx : Syntax) : Nat :=
match stx.getTailInfo? with
| some (SourceInfo.original (trailing := trailing) ..) => trailing.bsize
| _ => 0
/--
Return substring of original input covering `stx`.
Result is meaningful only if all involved `SourceInfo.original`s refer to the same string (as is the case after parsing). -/
def getSubstring? (stx : Syntax) (withLeading := true) (withTrailing := true) : Option Substring :=
match stx.getHeadInfo, stx.getTailInfo with
| SourceInfo.original lead startPos _ _, SourceInfo.original _ _ trail stopPos =>
some {
str := lead.str
startPos := if withLeading then lead.startPos else startPos
stopPos := if withTrailing then trail.stopPos else stopPos
}
| _, _ => none
@[specialize] private partial def updateLast {α} [Inhabited α] (a : Array α) (f : α → Option α) (i : Nat) : Option (Array α) :=
if i == 0 then
none
else
let i := i - 1
let v := a[i]!
match f v with
| some v => some <| a.set! i v
| none => updateLast a f i
partial def setTailInfoAux (info : SourceInfo) : Syntax → Option Syntax
| atom _ val => some <| atom info val
| ident _ rawVal val pre => some <| ident info rawVal val pre
| node info k args =>
match updateLast args (setTailInfoAux info) args.size with
| some args => some <| node info k args
| none => none
| _ => none
def setTailInfo (stx : Syntax) (info : SourceInfo) : Syntax :=
match setTailInfoAux info stx with
| some stx => stx
| none => stx
def unsetTrailing (stx : Syntax) : Syntax :=
match stx.getTailInfo with
| SourceInfo.original lead pos _ endPos => stx.setTailInfo (SourceInfo.original lead pos "".toSubstring endPos)
| _ => stx
@[specialize] private partial def updateFirst {α} [Inhabited α] (a : Array α) (f : α → Option α) (i : Nat) : Option (Array α) :=
if h : i < a.size then
let v := a[i]
match f v with
| some v => some <| a.set ⟨i, h⟩ v
| none => updateFirst a f (i+1)
else
none
partial def setHeadInfoAux (info : SourceInfo) : Syntax → Option Syntax
| atom _ val => some <| atom info val
| ident _ rawVal val pre => some <| ident info rawVal val pre
| node i k args =>
match updateFirst args (setHeadInfoAux info) 0 with
| some args => some <| node i k args
| _ => none
| _ => none
def setHeadInfo (stx : Syntax) (info : SourceInfo) : Syntax :=
match setHeadInfoAux info stx with
| some stx => stx
| none => stx
def setInfo (info : SourceInfo) : Syntax → Syntax
| atom _ val => atom info val
| ident _ rawVal val pre => ident info rawVal val pre
| node _ kind args => node info kind args
| missing => missing
/-- Return the first atom/identifier that has position information -/
partial def getHead? : Syntax → Option Syntax
| stx@(atom info ..) => info.getPos?.map fun _ => stx
| stx@(ident info ..) => info.getPos?.map fun _ => stx
| node SourceInfo.none _ args => args.findSome? getHead?
| stx@(node ..) => stx
| _ => none
def copyHeadTailInfoFrom (target source : Syntax) : Syntax :=
target.setHeadInfo source.getHeadInfo |>.setTailInfo source.getTailInfo
/-- Ensure head position is synthetic. The server regards syntax as "original" only if both head and tail info are `original`. -/
def mkSynthetic (stx : Syntax) : Syntax :=
stx.setHeadInfo (SourceInfo.fromRef stx)
end Syntax
/-- Use the head atom/identifier of the current `ref` as the `ref` -/
@[inline] def withHeadRefOnly {m : Type → Type} [Monad m] [MonadRef m] {α} (x : m α) : m α := do
match (← getRef).getHead? with
| none => x
| some ref => withRef ref x
/-- Syntax objects for a Lean module. -/
structure Module where
header : Syntax
commands : Array Syntax
/--
Expand macros in the given syntax.
A node with kind `k` is visited only if `p k` is true.
Note that the default value for `p` returns false for `by ...` nodes.
This is a "hack". The tactic framework abuses the macro system to implement extensible tactics.
For example, one can define
```lean
syntax "my_trivial" : tactic -- extensible tactic
macro_rules | `(tactic| my_trivial) => `(tactic| decide)
macro_rules | `(tactic| my_trivial) => `(tactic| assumption)
```
When the tactic evaluator finds the tactic `my_trivial`, it tries to evaluate the `macro_rule` expansions
until one "works", i.e., the macro expansion is evaluated without producing an exception.
We say this solution is a bit hackish because the term elaborator may invoke `expandMacros` with `(p := fun _ => true)`,
and expand the tactic macros as just macros. In the example above, `my_trivial` would be replaced with `assumption`,
`decide` would not be tried if `assumption` fails at tactic evaluation time.
We are considering two possible solutions for this issue:
1- A proper extensible tactic feature that does not rely on the macro system.
2- Typed macros that know the syntax categories they're working in. Then, we would be able to select which
syntatic categories are expanded by `expandMacros`.
-/
partial def expandMacros (stx : Syntax) (p : SyntaxNodeKind → Bool := fun k => k != `Lean.Parser.Term.byTactic) : MacroM Syntax :=
withRef stx do
match stx with
| .node info k args => do
if p k then
match (← expandMacro? stx) with
| some stxNew => expandMacros stxNew
| none => do
let args ← Macro.withIncRecDepth stx <| args.mapM expandMacros
return .node info k args
else
return stx
| stx => return stx
/-! # Helper functions for processing Syntax programmatically -/
/--
Create an identifier copying the position from `src`.
To refer to a specific constant, use `mkCIdentFrom` instead. -/
def mkIdentFrom (src : Syntax) (val : Name) (canonical := false) : Ident :=
⟨Syntax.ident (SourceInfo.fromRef src canonical) (toString val).toSubstring val []⟩
def mkIdentFromRef [Monad m] [MonadRef m] (val : Name) (canonical := false) : m Ident := do
return mkIdentFrom (← getRef) val canonical
/--
Create an identifier referring to a constant `c` copying the position from `src`.
This variant of `mkIdentFrom` makes sure that the identifier cannot accidentally
be captured. -/
def mkCIdentFrom (src : Syntax) (c : Name) (canonical := false) : Ident :=
-- Remark: We use the reserved macro scope to make sure there are no accidental collision with our frontend
let id := addMacroScope `_internal c reservedMacroScope
⟨Syntax.ident (SourceInfo.fromRef src canonical) (toString id).toSubstring id [.decl c []]⟩
def mkCIdentFromRef [Monad m] [MonadRef m] (c : Name) (canonical := false) : m Syntax := do
return mkCIdentFrom (← getRef) c canonical
def mkCIdent (c : Name) : Ident :=
mkCIdentFrom Syntax.missing c
@[export lean_mk_syntax_ident]
def mkIdent (val : Name) : Ident :=
⟨Syntax.ident SourceInfo.none (toString val).toSubstring val []⟩
@[inline] def mkGroupNode (args : Array Syntax := #[]) : Syntax :=
mkNode groupKind args
def mkSepArray (as : Array Syntax) (sep : Syntax) : Array Syntax := Id.run do
let mut i := 0
let mut r := #[]
for a in as do
if i > 0 then
r := r.push sep |>.push a
else
r := r.push a
i := i + 1
return r
def mkOptionalNode (arg : Option Syntax) : Syntax :=
match arg with
| some arg => mkNullNode #[arg]
| none => mkNullNode #[]
def mkHole (ref : Syntax) (canonical := false) : Syntax :=
mkNode `Lean.Parser.Term.hole #[mkAtomFrom ref "_" canonical]
namespace Syntax
def mkSep (a : Array Syntax) (sep : Syntax) : Syntax :=
mkNullNode <| mkSepArray a sep
def SepArray.ofElems {sep} (elems : Array Syntax) : SepArray sep :=
⟨mkSepArray elems (if sep.isEmpty then mkNullNode else mkAtom sep)⟩
def SepArray.ofElemsUsingRef [Monad m] [MonadRef m] {sep} (elems : Array Syntax) : m (SepArray sep) := do
let ref ← getRef;
return ⟨mkSepArray elems (if sep.isEmpty then mkNullNode else mkAtomFrom ref sep)⟩
instance : Coe (Array Syntax) (SepArray sep) where
coe := SepArray.ofElems
instance : Coe (TSyntaxArray k) (TSepArray k sep) where
coe a := ⟨mkSepArray a.raw (mkAtom sep)⟩
/-- Create syntax representing a Lean term application, but avoid degenerate empty applications. -/
def mkApp (fn : Term) : (args : TSyntaxArray `term) → Term
| #[] => fn
| args => ⟨mkNode `Lean.Parser.Term.app #[fn, mkNullNode args.raw]⟩
def mkCApp (fn : Name) (args : TSyntaxArray `term) : Term :=
mkApp (mkCIdent fn) args
def mkLit (kind : SyntaxNodeKind) (val : String) (info := SourceInfo.none) : TSyntax kind :=
let atom : Syntax := Syntax.atom info val
mkNode kind #[atom]
def mkStrLit (val : String) (info := SourceInfo.none) : StrLit :=
mkLit strLitKind (String.quote val) info
def mkNumLit (val : String) (info := SourceInfo.none) : NumLit :=
mkLit numLitKind val info
def mkScientificLit (val : String) (info := SourceInfo.none) : TSyntax scientificLitKind :=
mkLit scientificLitKind val info
def mkNameLit (val : String) (info := SourceInfo.none) : NameLit :=
mkLit nameLitKind val info
/-! Recall that we don't have special Syntax constructors for storing numeric and string atoms.
The idea is to have an extensible approach where embedded DSLs may have new kind of atoms and/or
different ways of representing them. So, our atoms contain just the parsed string.
The main Lean parser uses the kind `numLitKind` for storing natural numbers that can be encoded
in binary, octal, decimal and hexadecimal format. `isNatLit` implements a "decoder"
for Syntax objects representing these numerals. -/
private partial def decodeBinLitAux (s : String) (i : String.Pos) (val : Nat) : Option Nat :=
if s.atEnd i then some val
else
let c := s.get i
if c == '0' then decodeBinLitAux s (s.next i) (2*val)
else if c == '1' then decodeBinLitAux s (s.next i) (2*val + 1)
else none
private partial def decodeOctalLitAux (s : String) (i : String.Pos) (val : Nat) : Option Nat :=
if s.atEnd i then some val
else
let c := s.get i
if '0' ≤ c && c ≤ '7' then decodeOctalLitAux s (s.next i) (8*val + c.toNat - '0'.toNat)
else none
private def decodeHexDigit (s : String) (i : String.Pos) : Option (Nat × String.Pos) :=
let c := s.get i
let i := s.next i
if '0' ≤ c && c ≤ '9' then some (c.toNat - '0'.toNat, i)
else if 'a' ≤ c && c ≤ 'f' then some (10 + c.toNat - 'a'.toNat, i)
else if 'A' ≤ c && c ≤ 'F' then some (10 + c.toNat - 'A'.toNat, i)
else none
private partial def decodeHexLitAux (s : String) (i : String.Pos) (val : Nat) : Option Nat :=
if s.atEnd i then some val
else match decodeHexDigit s i with
| some (d, i) => decodeHexLitAux s i (16*val + d)
| none => none
private partial def decodeDecimalLitAux (s : String) (i : String.Pos) (val : Nat) : Option Nat :=
if s.atEnd i then some val
else
let c := s.get i
if '0' ≤ c && c ≤ '9' then decodeDecimalLitAux s (s.next i) (10*val + c.toNat - '0'.toNat)
else none
def decodeNatLitVal? (s : String) : Option Nat :=
let len := s.length
if len == 0 then none
else
let c := s.get 0
if c == '0' then
if len == 1 then some 0
else
let c := s.get ⟨1⟩
if c == 'x' || c == 'X' then decodeHexLitAux s ⟨2⟩ 0
else if c == 'b' || c == 'B' then decodeBinLitAux s ⟨2⟩ 0
else if c == 'o' || c == 'O' then decodeOctalLitAux s ⟨2⟩ 0
else if c.isDigit then decodeDecimalLitAux s 0 0
else none
else if c.isDigit then decodeDecimalLitAux s 0 0
else none
def isLit? (litKind : SyntaxNodeKind) (stx : Syntax) : Option String :=
match stx with
| Syntax.node _ k args =>
if k == litKind && args.size == 1 then
match args.get! 0 with
| (Syntax.atom _ val) => some val
| _ => none
else
none
| _ => none
private def isNatLitAux (litKind : SyntaxNodeKind) (stx : Syntax) : Option Nat :=
match isLit? litKind stx with
| some val => decodeNatLitVal? val
| _ => none
def isNatLit? (s : Syntax) : Option Nat :=
isNatLitAux numLitKind s
def isFieldIdx? (s : Syntax) : Option Nat :=
isNatLitAux fieldIdxKind s
/-- Decodes a 'scientific number' string which is consumed by the `OfScientific` class.
Takes as input a string such as `123`, `123.456e7` and returns a triple `(n, sign, e)` with value given by
`n * 10^-e` if `sign` else `n * 10^e`.
-/
partial def decodeScientificLitVal? (s : String) : Option (Nat × Bool × Nat) :=
let len := s.length
if len == 0 then none
else
let c := s.get 0
if c.isDigit then
decode 0 0
else none
where
decodeAfterExp (i : String.Pos) (val : Nat) (e : Nat) (sign : Bool) (exp : Nat) : Option (Nat × Bool × Nat) :=
if s.atEnd i then
if sign then
some (val, sign, exp + e)
else if exp >= e then
some (val, sign, exp - e)
else
some (val, true, e - exp)
else
let c := s.get i
if '0' ≤ c && c ≤ '9' then
decodeAfterExp (s.next i) val e sign (10*exp + c.toNat - '0'.toNat)
else
none
decodeExp (i : String.Pos) (val : Nat) (e : Nat) : Option (Nat × Bool × Nat) :=
if s.atEnd i then none else
let c := s.get i
if c == '-' then
decodeAfterExp (s.next i) val e true 0
else if c == '+' then
decodeAfterExp (s.next i) val e false 0
else
decodeAfterExp i val e false 0
decodeAfterDot (i : String.Pos) (val : Nat) (e : Nat) : Option (Nat × Bool × Nat) :=
if s.atEnd i then
some (val, true, e)
else
let c := s.get i
if '0' ≤ c && c ≤ '9' then
decodeAfterDot (s.next i) (10*val + c.toNat - '0'.toNat) (e+1)
else if c == 'e' || c == 'E' then
decodeExp (s.next i) val e
else
none
decode (i : String.Pos) (val : Nat) : Option (Nat × Bool × Nat) :=
if s.atEnd i then
none
else
let c := s.get i
if '0' ≤ c && c ≤ '9' then
decode (s.next i) (10*val + c.toNat - '0'.toNat)
else if c == '.' then
decodeAfterDot (s.next i) val 0
else if c == 'e' || c == 'E' then
decodeExp (s.next i) val 0
else
none
def isScientificLit? (stx : Syntax) : Option (Nat × Bool × Nat) :=
match isLit? scientificLitKind stx with
| some val => decodeScientificLitVal? val
| _ => none
def isIdOrAtom? : Syntax → Option String
| Syntax.atom _ val => some val
| Syntax.ident _ rawVal _ _ => some rawVal.toString
| _ => none
def toNat (stx : Syntax) : Nat :=
match stx.isNatLit? with
| some val => val
| none => 0
def decodeQuotedChar (s : String) (i : String.Pos) : Option (Char × String.Pos) := do
let c := s.get i
let i := s.next i
if c == '\\' then pure ('\\', i)
else if c = '\"' then pure ('\"', i)
else if c = '\'' then pure ('\'', i)
else if c = 'r' then pure ('\r', i)
else if c = 'n' then pure ('\n', i)
else if c = 't' then pure ('\t', i)
else if c = 'x' then
let (d₁, i) ← decodeHexDigit s i
let (d₂, i) ← decodeHexDigit s i
pure (Char.ofNat (16*d₁ + d₂), i)
else if c = 'u' then do
let (d₁, i) ← decodeHexDigit s i
let (d₂, i) ← decodeHexDigit s i
let (d₃, i) ← decodeHexDigit s i
let (d₄, i) ← decodeHexDigit s i
pure (Char.ofNat (16*(16*(16*d₁ + d₂) + d₃) + d₄), i)
else
none
partial def decodeStrLitAux (s : String) (i : String.Pos) (acc : String) : Option String := do
let c := s.get i
let i := s.next i
if c == '\"' then
pure acc
else if s.atEnd i then
none
else if c == '\\' then do
let (c, i) ← decodeQuotedChar s i
decodeStrLitAux s i (acc.push c)
else
decodeStrLitAux s i (acc.push c)
def decodeStrLit (s : String) : Option String :=
decodeStrLitAux s ⟨1⟩ ""
def isStrLit? (stx : Syntax) : Option String :=
match isLit? strLitKind stx with
| some val => decodeStrLit val
| _ => none
def decodeCharLit (s : String) : Option Char := do
let c := s.get ⟨1⟩
if c == '\\' then do
let (c, _) ← decodeQuotedChar s ⟨2⟩
pure c
else
pure c
def isCharLit? (stx : Syntax) : Option Char :=
match isLit? charLitKind stx with
| some val => decodeCharLit val
| _ => none
private partial def splitNameLitAux (ss : Substring) (acc : List Substring) : List Substring :=
let splitRest (ss : Substring) (acc : List Substring) : List Substring :=
if ss.front == '.' then
splitNameLitAux (ss.drop 1) acc
else if ss.isEmpty then
acc
else
[]
if ss.isEmpty then []
else
let curr := ss.front
if isIdBeginEscape curr then
let escapedPart := ss.takeWhile (!isIdEndEscape ·)
let escapedPart := { escapedPart with stopPos := ss.stopPos.min (escapedPart.str.next escapedPart.stopPos) }
if !isIdEndEscape (escapedPart.get <| escapedPart.prev ⟨escapedPart.bsize⟩) then []
else splitRest (ss.extract ⟨escapedPart.bsize⟩ ⟨ss.bsize⟩) (escapedPart :: acc)
else if isIdFirst curr then
let idPart := ss.takeWhile isIdRest
splitRest (ss.extract ⟨idPart.bsize⟩ ⟨ss.bsize⟩) (idPart :: acc)
else if curr.isDigit then
let idPart := ss.takeWhile Char.isDigit
splitRest (ss.extract ⟨idPart.bsize⟩ ⟨ss.bsize⟩) (idPart :: acc)
else
[]
/-- Split a name literal (without the backtick) into its dot-separated components. For example,
`foo.bla.«bo.o»` ↦ `["foo", "bla", "«bo.o»"]`. If the literal cannot be parsed, return `[]`. -/
def splitNameLit (ss : Substring) : List Substring :=
splitNameLitAux ss [] |>.reverse
def _root_.Substring.toName (s : Substring) : Name :=
match splitNameLitAux s [] with
| [] => .anonymous
| comps => comps.foldr (init := Name.anonymous)
fun comp n =>
let comp := comp.toString
if isIdBeginEscape comp.front then
Name.mkStr n (comp.drop 1 |>.dropRight 1)
else if comp.front.isDigit then
if let some k := decodeNatLitVal? comp then
Name.mkNum n k
else
unreachable!
else
Name.mkStr n comp
def _root_.String.toName (s : String) : Name :=
s.toSubstring.toName
def decodeNameLit (s : String) : Option Name :=
if s.get 0 == '`' then
match (s.toSubstring.drop 1).toName with
| .anonymous => none
| name => some name
else
none
def isNameLit? (stx : Syntax) : Option Name :=
match isLit? nameLitKind stx with
| some val => decodeNameLit val
| _ => none
def hasArgs : Syntax → Bool
| Syntax.node _ _ args => args.size > 0
| _ => false
def isAtom : Syntax → Bool
| atom _ _ => true
| _ => false
def isToken (token : String) : Syntax → Bool
| atom _ val => val.trim == token.trim
| _ => false
def isNone (stx : Syntax) : Bool :=
match stx with
| Syntax.node _ k args => k == nullKind && args.size == 0
-- when elaborating partial syntax trees, it's reasonable to interpret missing parts as `none`
| Syntax.missing => true
| _ => false
def getOptionalIdent? (stx : Syntax) : Option Name :=
match stx.getOptional? with
| some stx => some stx.getId
| none => none
partial def findAux (p : Syntax → Bool) : Syntax → Option Syntax
| stx@(Syntax.node _ _ args) => if p stx then some stx else args.findSome? (findAux p)
| stx => if p stx then some stx else none
def find? (stx : Syntax) (p : Syntax → Bool) : Option Syntax :=
findAux p stx
end Syntax
namespace TSyntax
def getNat (s : NumLit) : Nat :=
s.raw.isNatLit?.getD 0
def getId (s : Ident) : Name :=
s.raw.getId
def getScientific (s : ScientificLit) : Nat × Bool × Nat :=
s.raw.isScientificLit?.getD (0, false, 0)
def getString (s : StrLit) : String :=
s.raw.isStrLit?.getD ""
def getChar (s : CharLit) : Char :=
s.raw.isCharLit?.getD default
def getName (s : NameLit) : Name :=
s.raw.isNameLit?.getD .anonymous
def getHygieneInfo (s : HygieneInfo) : Name :=
s.raw[0].getId
namespace Compat
scoped instance : CoeTail (Array Syntax) (Syntax.TSepArray k sep) where
coe a := (a : TSyntaxArray k)
end Compat
end TSyntax
def HygieneInfo.mkIdent (s : HygieneInfo) (val : Name) (canonical := false) : Ident :=
let src := s.raw[0]
let id := { extractMacroScopes src.getId with name := val.eraseMacroScopes }.review
⟨Syntax.ident (SourceInfo.fromRef src canonical) (toString val).toSubstring id []⟩
/-- Reflect a runtime datum back to surface syntax (best-effort). -/
class Quote (α : Type) (k : SyntaxNodeKind := `term) where
quote : α → TSyntax k
export Quote (quote)
set_option synthInstance.checkSynthOrder false in
instance [Quote α k] [CoeHTCT (TSyntax k) (TSyntax [k'])] : Quote α k' := ⟨fun a => quote (k := k) a⟩
instance : Quote Term := ⟨id⟩
instance : Quote Bool := ⟨fun | true => mkCIdent ``Bool.true | false => mkCIdent ``Bool.false⟩
instance : Quote String strLitKind := ⟨Syntax.mkStrLit⟩
instance : Quote Nat numLitKind := ⟨fun n => Syntax.mkNumLit <| toString n⟩
instance : Quote Substring := ⟨fun s => Syntax.mkCApp ``String.toSubstring' #[quote s.toString]⟩
-- in contrast to `Name.toString`, we can, and want to be, precise here
private def getEscapedNameParts? (acc : List String) : Name → Option (List String)
| Name.anonymous => if acc.isEmpty then none else some acc
| Name.str n s => do
let s ← Name.escapePart s
getEscapedNameParts? (s::acc) n
| Name.num _ _ => none
def quoteNameMk : Name → Term
| .anonymous => mkCIdent ``Name.anonymous
| .str n s => Syntax.mkCApp ``Name.mkStr #[quoteNameMk n, quote s]
| .num n i => Syntax.mkCApp ``Name.mkNum #[quoteNameMk n, quote i]
instance : Quote Name `term where
quote n := match getEscapedNameParts? [] n with
| some ss => ⟨mkNode `Lean.Parser.Term.quotedName #[Syntax.mkNameLit ("`" ++ ".".intercalate ss)]⟩
| none => ⟨quoteNameMk n⟩
instance [Quote α `term] [Quote β `term] : Quote (α × β) `term where
quote
| ⟨a, b⟩ => Syntax.mkCApp ``Prod.mk #[quote a, quote b]
private def quoteList [Quote α `term] : List α → Term
| [] => mkCIdent ``List.nil
| (x::xs) => Syntax.mkCApp ``List.cons #[quote x, quoteList xs]
instance [Quote α `term] : Quote (List α) `term where
quote := quoteList
private def quoteArray [Quote α `term] (xs : Array α) : Term :=
if xs.size <= 8 then
go 0 #[]
else
Syntax.mkCApp ``List.toArray #[quote xs.toList]
where
go (i : Nat) (args : Array Term) : Term :=
if h : i < xs.size then
go (i+1) (args.push (quote xs[i]))
else
Syntax.mkCApp (Name.mkStr2 "Array" ("mkArray" ++ toString xs.size)) args
termination_by go i _ => xs.size - i
instance [Quote α `term] : Quote (Array α) `term where
quote := quoteArray