-
Notifications
You must be signed in to change notification settings - Fork 447
/
Copy pathApp.lean
1718 lines (1587 loc) · 82.1 KB
/
App.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
-/
prelude
import Lean.Util.FindMVar
import Lean.Util.CollectFVars
import Lean.Parser.Term
import Lean.Meta.KAbstract
import Lean.Meta.Tactic.ElimInfo
import Lean.Elab.Term
import Lean.Elab.Binders
import Lean.Elab.SyntheticMVars
import Lean.Elab.Arg
import Lean.Elab.RecAppSyntax
namespace Lean.Elab.Term
open Meta
builtin_initialize elabWithoutExpectedTypeAttr : TagAttribute ←
registerTagAttribute `elab_without_expected_type "mark that applications of the given declaration should be elaborated without the expected type"
def hasElabWithoutExpectedType (env : Environment) (declName : Name) : Bool :=
elabWithoutExpectedTypeAttr.hasTag env declName
instance : ToString Arg where
toString
| .stx val => toString val
| .expr val => toString val
instance : ToString NamedArg where
toString s := "(" ++ toString s.name ++ " := " ++ toString s.val ++ ")"
def throwInvalidNamedArg (namedArg : NamedArg) (fn? : Option Name) : TermElabM α :=
withRef namedArg.ref <| match fn? with
| some fn => throwError "invalid argument name '{namedArg.name}' for function '{fn}'"
| none => throwError "invalid argument name '{namedArg.name}' for function"
private def ensureArgType (f : Expr) (arg : Expr) (expectedType : Expr) : TermElabM Expr := do
try
ensureHasType expectedType arg none f
catch
| ex@(.error ..) =>
if (← read).errToSorry then
exceptionToSorry ex expectedType
else
throw ex
| ex => throw ex
private def mkProjAndCheck (structName : Name) (idx : Nat) (e : Expr) : MetaM Expr := do
let r := mkProj structName idx e
let eType ← inferType e
if (← isProp eType) then
let rType ← inferType r
if !(← isProp rType) then
throwError "invalid projection, the expression{indentExpr e}\nis a proposition and has type{indentExpr eType}\nbut the projected value is not, it has type{indentExpr rType}"
return r
def synthesizeAppInstMVars (instMVars : Array MVarId) (app : Expr) : TermElabM Unit :=
for mvarId in instMVars do
unless (← synthesizeInstMVarCore mvarId) do
registerSyntheticMVarWithCurrRef mvarId (.typeClass none)
registerMVarErrorImplicitArgInfo mvarId (← getRef) app
/-- Return `some namedArg` if `namedArgs` contains an entry for `binderName`. -/
private def findBinderName? (namedArgs : List NamedArg) (binderName : Name) : Option NamedArg :=
namedArgs.find? fun namedArg => namedArg.name == binderName
/-- Erase entry for `binderName` from `namedArgs`. -/
def eraseNamedArg (namedArgs : List NamedArg) (binderName : Name) : List NamedArg :=
namedArgs.filter (·.name != binderName)
/-- Return true if the given type contains `OptParam` or `AutoParams` -/
private def hasOptAutoParams (type : Expr) : MetaM Bool := do
forallTelescopeReducing type fun xs _ =>
xs.anyM fun x => do
let xType ← inferType x
return xType.getOptParamDefault?.isSome || xType.getAutoParamTactic?.isSome
/-! # Default application elaborator -/
namespace ElabAppArgs
structure Context where
/--
`true` if `..` was used
-/
ellipsis : Bool --
/--
`true` if `@` modifier was used
-/
explicit : Bool
/--
If the result type of an application is the `outParam` of some local instance, then special support may be needed
because type class resolution interacts poorly with coercions in this kind of situation.
This flag enables the special support.
The idea is quite simple, if the result type is the `outParam` of some local instance, we simply
execute `synthesizeSyntheticMVarsUsingDefault`. We added this feature to make sure examples as follows
are correctly elaborated.
```lean
class GetElem (Cont : Type u) (Idx : Type v) (Elem : outParam (Type w)) where
getElem (xs : Cont) (i : Idx) : Elem
export GetElem (getElem)
instance : GetElem (Array α) Nat α where
getElem xs i := xs.get ⟨i, sorry⟩
opaque f : Option Bool → Bool
opaque g : Bool → Bool
def bad (xs : Array Bool) : Bool :=
let x := getElem xs 0
f x && g x
```
Without the special support, Lean fails at `g x` saying `x` has type `Option Bool` but is expected to have type `Bool`.
From the user's point of view this is a bug, since `let x := getElem xs 0` clearly constrains `x` to be `Bool`, but
we only obtain this information after we apply the `OfNat` default instance for `0`.
Before converging to this solution, we have tried to create a "coercion placeholder" when `resultIsOutParamSupport = true`,
but it did not work well in practice. For example, it failed in the example above.
-/
resultIsOutParamSupport : Bool
/-- Auxiliary structure for elaborating the application `f args namedArgs`. -/
structure State where
f : Expr
fType : Expr
/-- Remaining regular arguments. -/
args : List Arg
/-- remaining named arguments to be processed. -/
namedArgs : List NamedArg
expectedType? : Option Expr
/--
When named arguments are provided and explicit arguments occurring before them are missing,
the elaborator eta-expands the declaration. For example,
```
def f (x y : Nat) := x + y
#check f (y := 5)
-- fun x => f x 5
```
`etaArgs` stores the fresh free variables for implementing the eta-expansion.
When `..` is used, eta-expansion is disabled, and missing arguments are treated as `_`.
-/
etaArgs : Array Expr := #[]
/-- Metavariables that we need to set the error context using the application being built. -/
toSetErrorCtx : Array MVarId := #[]
/-- Metavariables for the instance implicit arguments that have already been processed. -/
instMVars : Array MVarId := #[]
/--
The following field is used to implement the `propagateExpectedType` heuristic.
It is set to `true` true when `expectedType` still has to be propagated.
-/
propagateExpected : Bool
/--
If the result type may be the `outParam` of some local instance.
See comment at `Context.resultIsOutParamSupport`
-/
resultTypeOutParam? : Option MVarId := none
abbrev M := ReaderT Context (StateRefT State TermElabM)
/-- Add the given metavariable to the collection of metavariables associated with instance-implicit arguments. -/
private def addInstMVar (mvarId : MVarId) : M Unit :=
modify fun s => { s with instMVars := s.instMVars.push mvarId }
/--
Try to synthesize metavariables are `instMVars` using type class resolution.
The ones that cannot be synthesized yet stay in the `instMVars` list.
Remark: we use this method
- before trying to apply coercions to function,
- before unifying the expected type.
-/
def trySynthesizeAppInstMVars : M Unit := do
let instMVars ← (← get).instMVars.filterM fun instMVar => do
unless (← instantiateMVars (← inferType (.mvar instMVar))).isMVar do try
if (← synthesizeInstMVarCore instMVar) then
return false
catch _ => pure ()
return true
modify ({ · with instMVars })
/--
Try to synthesize metavariables are `instMVars` using type class resolution.
The ones that cannot be synthesized yet are registered.
-/
def synthesizeAppInstMVars : M Unit := do
Term.synthesizeAppInstMVars (← get).instMVars (← get).f
modify ({ · with instMVars := #[] })
/-- fType may become a forallE after we synthesize pending metavariables. -/
private def synthesizePendingAndNormalizeFunType : M Unit := do
trySynthesizeAppInstMVars
synthesizeSyntheticMVars
let s ← get
let fType ← whnfForall s.fType
if fType.isForall then
modify fun s => { s with fType }
else
if let some f ← coerceToFunction? s.f then
let fType ← inferType f
modify fun s => { s with f, fType }
else
for namedArg in s.namedArgs do
let f := s.f.getAppFn
if f.isConst then
throwInvalidNamedArg namedArg f.constName!
else
throwInvalidNamedArg namedArg none
throwError "function expected at{indentExpr s.f}\nterm has type{indentExpr fType}"
/-- Normalize and return the function type. -/
private def normalizeFunType : M Expr := do
let s ← get
let fType ← whnfForall s.fType
modify fun s => { s with fType }
return fType
/-- Return the binder name at `fType`. This method assumes `fType` is a function type. -/
private def getBindingName : M Name := return (← get).fType.bindingName!
/-- Return the next argument expected type. This method assumes `fType` is a function type. -/
private def getArgExpectedType : M Expr := return (← get).fType.bindingDomain!
/-- Remove named argument with name `binderName` from `namedArgs`. -/
def eraseNamedArg (binderName : Name) : M Unit :=
modify fun s => { s with namedArgs := Term.eraseNamedArg s.namedArgs binderName }
/--
Add a new argument to the result. That is, `f := f arg`, update `fType`.
This method assumes `fType` is a function type. -/
private def addNewArg (argName : Name) (arg : Expr) : M Unit := do
modify fun s => { s with f := mkApp s.f arg, fType := s.fType.bindingBody!.instantiate1 arg }
if arg.isMVar then
registerMVarArgName arg.mvarId! argName
/--
Elaborate the given `Arg` and add it to the result. See `addNewArg`.
Recall that, `Arg` may be wrapping an already elaborated `Expr`. -/
private def elabAndAddNewArg (argName : Name) (arg : Arg) : M Unit := do
let s ← get
let expectedType := (← getArgExpectedType).consumeTypeAnnotations
match arg with
| Arg.expr val =>
let arg ← ensureArgType s.f val expectedType
addNewArg argName arg
| Arg.stx stx =>
let val ← elabTerm stx expectedType
let arg ← withRef stx <| ensureArgType s.f val expectedType
addNewArg argName arg
/-- Return true if `fType` contains `OptParam` or `AutoParams` -/
private def fTypeHasOptAutoParams : M Bool := do
hasOptAutoParams (← get).fType
/--
Auxiliary function for retrieving the resulting type of a function application.
See `propagateExpectedType`.
Remark: `(explicit : Bool) == true` when `@` modifier is used. -/
private partial def getForallBody (explicit : Bool) : Nat → List NamedArg → Expr → Option Expr
| i, namedArgs, type@(.forallE n d b bi) =>
match findBinderName? namedArgs n with
| some _ => getForallBody explicit i (Term.eraseNamedArg namedArgs n) b
| none =>
if !explicit && !bi.isExplicit then
getForallBody explicit i namedArgs b
else if i > 0 then
getForallBody explicit (i-1) namedArgs b
else if d.isAutoParam || d.isOptParam then
getForallBody explicit i namedArgs b
else
some type
| 0, [], type => some type
| _, _, _ => none
private def shouldPropagateExpectedTypeFor (nextArg : Arg) : Bool :=
match nextArg with
| .expr _ => false -- it has already been elaborated
| .stx stx =>
-- TODO: make this configurable?
stx.getKind != ``Lean.Parser.Term.hole &&
stx.getKind != ``Lean.Parser.Term.syntheticHole &&
stx.getKind != ``Lean.Parser.Term.byTactic
/--
Auxiliary method for propagating the expected type. We call it as soon as we find the first explicit
argument. The goal is to propagate the expected type in applications of functions such as
```lean
Add.add {α : Type u} : α → α → α
List.cons {α : Type u} : α → List α → List α
```
This is particularly useful when there applicable coercions. For example,
assume we have a coercion from `Nat` to `Int`, and we have
`(x : Nat)` and the expected type is `List Int`. Then, if we don't use this function,
the elaborator will fail to elaborate
```
List.cons x []
```
First, the elaborator creates a new metavariable `?α` for the implicit argument `{α : Type u}`.
Then, when it processes `x`, it assigns `?α := Nat`, and then obtains the
resultant type `List Nat` which is **not** definitionally equal to `List Int`.
We solve the problem by executing this method before we elaborate the first explicit argument (`x` in this example).
This method infers that the resultant type is `List ?α` and unifies it with `List Int`.
Then, when we elaborate `x`, the elaborate realizes the coercion from `Nat` to `Int` must be used, and the
term
```
@List.cons Int (coe x) (@List.nil Int)
```
is produced.
The method will do nothing if
1- The resultant type depends on the remaining arguments (i.e., `!eTypeBody.hasLooseBVars`).
2- The resultant type contains optional/auto params.
We have considered adding the following extra conditions
a) The resultant type does not contain any type metavariable.
b) The resultant type contains a nontype metavariable.
These two conditions would restrict the method to simple functions that are "morally" in
the Hindley&Milner fragment.
If users need to disable expected type propagation, we can add an attribute `[elab_without_expected_type]`.
-/
private def propagateExpectedType (arg : Arg) : M Unit := do
if shouldPropagateExpectedTypeFor arg then
let s ← get
-- TODO: handle s.etaArgs.size > 0
unless !s.etaArgs.isEmpty || !s.propagateExpected do
match s.expectedType? with
| none => pure ()
| some expectedType =>
/- We don't propagate `Prop` because we often use `Prop` as a more general "Bool" (e.g., `if-then-else`).
If we propagate `expectedType == Prop` in the following examples, the elaborator would fail
```
def f1 (s : Nat × Bool) : Bool := if s.2 then false else true
def f2 (s : List Bool) : Bool := if s.head! then false else true
def f3 (s : List Bool) : Bool := if List.head! (s.map not) then false else true
```
They would all fail for the same reason. So, let's focus on the first one.
We would elaborate `s.2` with `expectedType == Prop`.
Before we elaborate `s`, this method would be invoked, and `s.fType` is `?α × ?β → ?β` and after
propagation we would have `?α × Prop → Prop`. Then, when we would try to elaborate `s`, and
get a type error because `?α × Prop` cannot be unified with `Nat × Bool`.
Most users would have a hard time trying to understand why these examples failed.
Here is a possible alternative workaround. We give up the idea of using `Prop` at `if-then-else`.
Drawback: users use `if-then-else` with conditions that are not Decidable.
So, users would have to embrace `propDecidable` and `choice`.
This may not be that bad since the developers and users don't seem to care about constructivism.
We currently use a different workaround, we just don't propagate the expected type when it is `Prop`. -/
if expectedType.isProp then
modify fun s => { s with propagateExpected := false }
else
let numRemainingArgs := s.args.length
trace[Elab.app.propagateExpectedType] "etaArgs.size: {s.etaArgs.size}, numRemainingArgs: {numRemainingArgs}, fType: {s.fType}"
match getForallBody (← read).explicit numRemainingArgs s.namedArgs s.fType with
| none => pure ()
| some fTypeBody =>
unless fTypeBody.hasLooseBVars do
unless (← hasOptAutoParams fTypeBody) do
trySynthesizeAppInstMVars
trace[Elab.app.propagateExpectedType] "{expectedType} =?= {fTypeBody}"
if (← isDefEq expectedType fTypeBody) then
/- Note that we only set `propagateExpected := false` when propagation has succeeded. -/
modify fun s => { s with propagateExpected := false }
/-- This method executes after all application arguments have been processed. -/
private def finalize : M Expr := do
let s ← get
let mut e := s.f
-- all user explicit arguments have been consumed
trace[Elab.app.finalize] e
let ref ← getRef
-- Register the error context of implicits
for mvarId in s.toSetErrorCtx do
registerMVarErrorImplicitArgInfo mvarId ref e
if !s.etaArgs.isEmpty then
e ← mkLambdaFVars s.etaArgs e
/-
Remark: we should not use `s.fType` as `eType` even when
`s.etaArgs.isEmpty`. Reason: it may have been unfolded.
-/
let eType ← inferType e
trace[Elab.app.finalize] "after etaArgs, {e} : {eType}"
/- Recall that `resultTypeOutParam? = some mvarId` if the function result type is the output parameter
of a local instance. The value of this parameter may be inferable using other arguments. For example,
suppose we have
```lean
def add_one {X} [Trait X] [One (Trait.R X)] [HAdd X (Trait.R X) X] (x : X) : X := x + (One.one : (Trait.R X))
```
from test `948.lean`. There are multiple ways to infer `X`, and we don't want to mark it as `syntheticOpaque`.
-/
if let some outParamMVarId := s.resultTypeOutParam? then
synthesizeAppInstMVars
/- If `eType != mkMVar outParamMVarId`, then the
function is partially applied, and we do not apply default instances. -/
if !(← outParamMVarId.isAssigned) && eType.isMVar && eType.mvarId! == outParamMVarId then
synthesizeSyntheticMVarsUsingDefault
return e
else
return e
if let some expectedType := s.expectedType? then
trySynthesizeAppInstMVars
-- Try to propagate expected type. Ignore if types are not definitionally equal, caller must handle it.
trace[Elab.app.finalize] "expected type: {expectedType}"
discard <| isDefEq expectedType eType
synthesizeAppInstMVars
return e
/--
Returns a named argument that depends on the next argument, otherwise `none`.
-/
private def findNamedArgDependsOnCurrent? : M (Option NamedArg) := do
let s ← get
if s.namedArgs.isEmpty then
return none
else
forallTelescopeReducing s.fType fun xs _ => do
let curr := xs[0]!
for h : i in [1:xs.size] do
let xDecl ← xs[i].fvarId!.getDecl
if let some arg := s.namedArgs.find? fun arg => arg.name == xDecl.userName then
/- Remark: a default value at `optParam` does not count as a dependency -/
if (← exprDependsOn xDecl.type.cleanupAnnotations curr.fvarId!) then
return arg
return none
/-- Return `true` if there are regular or named arguments to be processed. -/
private def hasArgsToProcess : M Bool := do
let s ← get
return !s.args.isEmpty || !s.namedArgs.isEmpty
/--
Returns the argument syntax if the next argument at `args` is of the form `_`.
-/
private def nextArgHole? : M (Option Syntax) := do
match (← get).args with
| Arg.stx stx@(Syntax.node _ ``Lean.Parser.Term.hole _) :: _ => pure stx
| _ => pure none
/--
Return `true` if the next argument to be processed is the outparam of a local instance, and it the result type
of the function.
For example, suppose we have the class
```lean
class Get (Cont : Type u) (Idx : Type v) (Elem : outParam (Type w)) where
get (xs : Cont) (i : Idx) : Elem
```
And the current value of `fType` is
```
{Cont : Type u_1} → {Idx : Type u_2} → {Elem : Type u_3} → [self : Get Cont Idx Elem] → Cont → Idx → Elem
```
then the result returned by this method is `false` since `Cont` is not the output param of any local instance.
Now assume `fType` is
```
{Elem : Type u_3} → [self : Get Cont Idx Elem] → Cont → Idx → Elem
```
then, the method returns `true` because `Elem` is an output parameter for the local instance `[self : Get Cont Idx Elem]`.
Remark: if `resultIsOutParamSupport` is `false`, this method returns `false`.
-/
private partial def isNextOutParamOfLocalInstanceAndResult : M Bool := do
if !(← read).resultIsOutParamSupport then
return false
let type := (← get).fType.bindingBody!
unless isResultType type 0 do
return false
if (← hasLocalInstaceWithOutParams type) then
let x := mkFVar (← mkFreshFVarId)
isOutParamOfLocalInstance x (type.instantiate1 x)
else
return false
where
isResultType (type : Expr) (i : Nat) : Bool :=
match type with
| .forallE _ _ b _ => isResultType b (i + 1)
| .bvar idx => idx == i
| _ => false
/-- (quick filter) Return true if `type` contains a binder `[C ...]` where `C` is a class containing outparams. -/
hasLocalInstaceWithOutParams (type : Expr) : CoreM Bool := do
let .forallE _ d b bi := type | return false
if bi.isInstImplicit then
if let .const declName .. := d.getAppFn then
if hasOutParams (← getEnv) declName then
return true
hasLocalInstaceWithOutParams b
isOutParamOfLocalInstance (x : Expr) (type : Expr) : MetaM Bool := do
let .forallE _ d b bi := type | return false
if bi.isInstImplicit then
if let .const declName .. := d.getAppFn then
if hasOutParams (← getEnv) declName then
let cType ← inferType d.getAppFn
if (← isOutParamOf x 0 d.getAppArgs cType) then
return true
isOutParamOfLocalInstance x b
isOutParamOf (x : Expr) (i : Nat) (args : Array Expr) (cType : Expr) : MetaM Bool := do
if h : i < args.size then
match (← whnf cType) with
| .forallE _ d b _ =>
if args[i] == x && d.isOutParam then
return true
isOutParamOf x (i+1) args b
| _ => return false
else
return false
mutual
/--
Create a fresh local variable with the current binder name and argument type, add it to `etaArgs` and `f`,
and then execute the main loop.
-/
private partial def addEtaArg (argName : Name) : M Expr := do
let n ← getBindingName
let type ← getArgExpectedType
withLocalDeclD n type fun x => do
modify fun s => { s with etaArgs := s.etaArgs.push x }
addNewArg argName x
main
/--
Create a fresh metavariable for the implicit argument, add it to `f`, and then execute the main loop.
-/
private partial def addImplicitArg (argName : Name) : M Expr := do
let argType ← getArgExpectedType
let arg ← if (← isNextOutParamOfLocalInstanceAndResult) then
let arg ← mkFreshExprMVar argType
/- When the result type is an output parameter, we don't want to propagate the expected type.
So, we just mark `propagateExpected := false` to disable it.
At `finalize`, we check whether `arg` is still unassigned, if it is, we apply default instances,
and try to synthesize pending mvars. -/
modify fun s => { s with resultTypeOutParam? := some arg.mvarId!, propagateExpected := false }
pure arg
else
mkFreshExprMVar argType
modify fun s => { s with toSetErrorCtx := s.toSetErrorCtx.push arg.mvarId! }
addNewArg argName arg
main
/--
Process a `fType` of the form `(x : A) → B x`.
This method assume `fType` is a function type.
-/
private partial def processExplicitArg (argName : Name) : M Expr := do
match (← get).args with
| arg::args =>
-- Note: currently the following test never succeeds in explicit mode since `@x.f` notation does not exist.
if let some true := NamedArg.suppressDeps <$> (← findNamedArgDependsOnCurrent?) then
/-
We treat the explicit argument `argName` as implicit
if we have a named arguments that depends on it whose `suppressDeps` flag set to `true`.
The motivation for this is class projections (issue #1851).
In some cases, class projections can have explicit parameters. For example, in
```
class Approx {α : Type} (a : α) (X : Type) : Type where
val : X
```
the type of `Approx.val` is `{α : Type} → (a : α) → {X : Type} → [self : Approx a X] → X`.
Note that the parameter `a` is explicit since there is no way to infer it from the expected
type or from the types of other explicit parameters.
Being a parameter of the class, `a` is determined by the type of `self`.
Consider
```
variable {α β X Y : Type} {f' : α → β} {x' : α} [f : Approx f' (X → Y)]
```
Recall that `f.val` is, to first approximation, sugar for `Approx.val (self := f)`.
Without further refinement, this would expand to `fun f'' : α → β => Approx.val f'' f`,
which is a type error, since `f''` must be defeq to `f'`.
Furthermore, with projection notation, users expect all structure parameters
to be uniformly implicit; after all, they are determined by `self`.
To handle this, the `(self := f)` named argument is annotated with the `suppressDeps` flag.
This causes the `a` parameter to become implicit, and `f.val` instead expands to `Approx.val f' f`.
This feature previously was enabled for *all* explicit arguments, which confused users
and was frequently reported as a bug (issue #1867).
Now it is only enabled for the `self` argument in structure projections.
We used to do this only when `(← get).args` was empty,
but it created an asymmetry because `f.val` worked as expected,
yet one would have to write `f.val _ x` when there are further arguments.
-/
return (← addImplicitArg argName)
propagateExpectedType arg
modify fun s => { s with args }
elabAndAddNewArg argName arg
main
| _ =>
if (← read).ellipsis && (← readThe Term.Context).inPattern then
/-
In patterns, ellipsis should always be an implicit argument, even if it is an optparam or autoparam.
This prevents examples such as the one in #4555 from failing:
```lean
match e with
| .internal .. => sorry
| .error .. => sorry
```
The `internal` has an optparam (`| internal (id : InternalExceptionId) (extra : KVMap := {})`).
We may consider having ellipsis suppress optparams and autoparams in general.
We avoid doing so for now since it's possible to opt-out of them (for example with `.internal (extra := _) ..`)
but it's not possible to opt-in.
-/
return ← addImplicitArg argName
let argType ← getArgExpectedType
match (← read).explicit, argType.getOptParamDefault?, argType.getAutoParamTactic? with
| false, some defVal, _ => addNewArg argName defVal; main
| false, _, some (.const tacticDecl _) =>
let env ← getEnv
let opts ← getOptions
match evalSyntaxConstant env opts tacticDecl with
| Except.error err => throwError err
| Except.ok tacticSyntax =>
let tacticBlock ← `(by $(⟨tacticSyntax⟩))
/-
We insert position information from the current ref into `stx` everywhere, simulating this being
a tactic script inserted by the user, which ensures error messages and logging will always be attributed
to this application rather than sometimes being placed at position (1,0) in the file.
Placing position information on `by` syntax alone is not sufficient since incrementality
(in particular, `Lean.Elab.Term.withReuseContext`) controls the ref to avoid leakage of outside data.
Note that `tacticSyntax` contains no position information itself, since it is erased by `Lean.Elab.Term.quoteAutoTactic`.
-/
let info := (← getRef).getHeadInfo
let tacticBlock := tacticBlock.raw.rewriteBottomUp (·.setInfo info)
let mvar ← mkTacticMVar argType.consumeTypeAnnotations tacticBlock (.autoParam argName)
-- Note(kmill): We are adding terminfo to simulate a previous implementation that elaborated `tacticBlock`.
-- We should look into removing this since terminfo for synthetic syntax is suspect,
-- but we noted it was necessary to preserve the behavior of the unused variable linter.
addTermInfo' tacticBlock mvar
let argNew := Arg.expr mvar
propagateExpectedType argNew
elabAndAddNewArg argName argNew
main
| false, _, some _ =>
throwError "invalid autoParam, argument must be a constant"
| _, _, _ =>
if (← read).ellipsis then
addImplicitArg argName
else if !(← get).namedArgs.isEmpty then
if let some _ ← findNamedArgDependsOnCurrent? then
/-
Dependencies of named arguments cannot be turned into eta arguments
since they are determined by the named arguments.
Instead we can turn them into implicit arguments.
-/
addImplicitArg argName
else
addEtaArg argName
else if !(← read).explicit then
if (← fTypeHasOptAutoParams) then
addEtaArg argName
else
finalize
else
finalize
/--
Process a `fType` of the form `{x : A} → B x`.
This method assume `fType` is a function type -/
private partial def processImplicitArg (argName : Name) : M Expr := do
if (← read).explicit then
processExplicitArg argName
else
addImplicitArg argName
/--
Process a `fType` of the form `{{x : A}} → B x`.
This method assume `fType` is a function type -/
private partial def processStrictImplicitArg (argName : Name) : M Expr := do
if (← read).explicit then
processExplicitArg argName
else if (← hasArgsToProcess) then
addImplicitArg argName
else
finalize
/--
Process a `fType` of the form `[x : A] → B x`.
This method assume `fType` is a function type.
-/
private partial def processInstImplicitArg (argName : Name) : M Expr := do
if (← read).explicit then
if let some stx ← nextArgHole? then
-- We still use typeclass resolution for `_` arguments.
-- This behavior can be suppressed with `(_)`.
let ty ← getArgExpectedType
let arg ← mkInstMVar ty
addTermInfo' stx arg ty
modify fun s => { s with args := s.args.tail! }
main
else
processExplicitArg argName
else
discard <| mkInstMVar (← getArgExpectedType)
main
where
mkInstMVar (ty : Expr) : M Expr := do
let arg ← mkFreshExprMVar ty MetavarKind.synthetic
addInstMVar arg.mvarId!
addNewArg argName arg
return arg
/-- Elaborate function application arguments. -/
partial def main : M Expr := do
let fType ← normalizeFunType
if fType.isForall then
let binderName := fType.bindingName!
let binfo := fType.bindingInfo!
let s ← get
match findBinderName? s.namedArgs binderName with
| some namedArg =>
propagateExpectedType namedArg.val
eraseNamedArg binderName
elabAndAddNewArg binderName namedArg.val
main
| none =>
match binfo with
| .implicit => processImplicitArg binderName
| .instImplicit => processInstImplicitArg binderName
| .strictImplicit => processStrictImplicitArg binderName
| _ => processExplicitArg binderName
else if (← hasArgsToProcess) then
synthesizePendingAndNormalizeFunType
main
else
finalize
end
end ElabAppArgs
/-! # Eliminator-like function application elaborator -/
/--
Information about an eliminator used by the elab-as-elim elaborator.
This is not to be confused with `Lean.Meta.ElimInfo`, which is for `induction` and `cases`.
The elab-as-elim routine is less restrictive in what counts as an eliminator, and it doesn't need
to have a strict notion of what is a "target" — all it cares about are
1. that the return type of a function is of the form `m ...` where `m` is a parameter
(unlike `induction` and `cases` eliminators, the arguments to `m`, known as "discriminants",
can be any expressions, not just parameters), and
2. which arguments should be eagerly elaborated, to make discriminants be as elaborated as
possible for the expected type generalization procedure,
and which should be postponed (since they are the "minor premises").
Note that the routine isn't doing induction/cases *on* particular expressions.
The purpose of elab-as-elim is to successfully solve the higher-order unification problem
between the return type of the function and the expected type.
-/
structure ElabElimInfo where
/-- The eliminator. -/
elimExpr : Expr
/-- The type of the eliminator. -/
elimType : Expr
/-- The position of the motive parameter. -/
motivePos : Nat
/--
Positions of "major" parameters (those that should be eagerly elaborated
because they can contribute to the motive inference procedure).
All parameters that are neither the motive nor a major parameter are "minor" parameters.
The major parameters include all of the parameters that transitively appear in the motive's arguments,
as well as "first-order" arguments that include such parameters,
since they too can help with elaborating discriminants.
For example, in the following theorem the argument `h : a = b`
should be elaborated eagerly because it contains `b`, which occurs in `motive b`.
```
theorem Eq.subst' {α} {motive : α → Prop} {a b : α} (h : a = b) : motive a → motive b
```
For another example, the term `isEmptyElim (α := α)` is an underapplied eliminator, and it needs
argument `α` to be elaborated eagerly to create a type-correct motive.
```
def isEmptyElim [IsEmpty α] {p : α → Sort _} (a : α) : p a := ...
example {α : Type _} [IsEmpty α] : id (α → False) := isEmptyElim (α := α)
```
-/
majorsPos : Array Nat := #[]
deriving Repr, Inhabited
def getElabElimExprInfo (elimExpr : Expr) : MetaM ElabElimInfo := do
let elimType ← inferType elimExpr
trace[Elab.app.elab_as_elim] "eliminator {indentExpr elimExpr}\nhas type{indentExpr elimType}"
forallTelescopeReducing elimType fun xs type => do
let motive := type.getAppFn
let motiveArgs := type.getAppArgs
unless motive.isFVar && motiveArgs.size > 0 do
throwError "unexpected eliminator resulting type{indentExpr type}"
let motiveType ← inferType motive
forallTelescopeReducing motiveType fun motiveParams motiveResultType => do
unless motiveParams.size == motiveArgs.size do
throwError "unexpected number of arguments at motive type{indentExpr motiveType}"
unless motiveResultType.isSort do
throwError "motive result type must be a sort{indentExpr motiveType}"
let some motivePos ← pure (xs.indexOf? motive) |
throwError "unexpected eliminator type{indentExpr elimType}"
/-
Compute transitive closure of fvars appearing in arguments to the motive.
These are the primary set of major parameters.
-/
let initMotiveFVars : CollectFVars.State := motiveArgs.foldl (init := {}) collectFVars
let motiveFVars ← xs.size.foldRevM (init := initMotiveFVars) fun i _ s => do
let x := xs[i]
if s.fvarSet.contains x.fvarId! then
return collectFVars s (← inferType x)
else
return s
/- Collect the major parameter positions -/
let mut majorsPos := #[]
for h : i in [:xs.size] do
let x := xs[i]
unless motivePos == i do
let xType ← x.fvarId!.getType
/-
We also consider "first-order" types because we can reliably "extract" information from them.
We say a term is "first-order" if all applications are of the form `f ...` where `f` is a constant.
-/
let isFirstOrder (e : Expr) : Bool := Option.isNone <| e.find? fun e => e.isApp && !e.getAppFn.isConst
if motiveFVars.fvarSet.contains x.fvarId!
|| (isFirstOrder xType
&& Option.isSome (xType.find? fun e => e.isFVar && motiveFVars.fvarSet.contains e.fvarId!)) then
majorsPos := majorsPos.push i
trace[Elab.app.elab_as_elim] "motivePos: {motivePos}"
trace[Elab.app.elab_as_elim] "majorsPos: {majorsPos}"
return { elimExpr, elimType, motivePos, majorsPos }
def getElabElimInfo (elimName : Name) : MetaM ElabElimInfo := do
getElabElimExprInfo (← mkConstWithFreshMVarLevels elimName)
builtin_initialize elabAsElim : TagAttribute ←
registerTagAttribute `elab_as_elim
"instructs elaborator that the arguments of the function application should be elaborated as were an eliminator"
/-
We apply `elab_as_elim` after compilation because this kind of attribute is not applied to auxiliary declarations
created by the `WF` and `Structural` modules. This is an "indirect" fix for issue #1900. We should consider
having an explicit flag in attributes to indicate whether they should be copied to auxiliary declarations or not.
-/
(applicationTime := .afterCompilation)
fun declName => do
let go : MetaM Unit := do
let info ← getConstInfo declName
if (← hasOptAutoParams info.type) then
throwError "[elab_as_elim] attribute cannot be used in declarations containing optional and auto parameters"
discard <| getElabElimInfo declName
go.run' {} {}
namespace ElabElim
/-- Context of the `elab_as_elim` elaboration procedure. -/
structure Context where
elimInfo : ElabElimInfo
expectedType : Expr
/-- State of the `elab_as_elim` elaboration procedure. -/
structure State where
/-- The resultant expression being built. -/
f : Expr
/-- `f : fType -/
fType : Expr
/-- User-provided named arguments that still have to be processed. -/
namedArgs : List NamedArg
/-- User-provided arguments that still have to be processed. -/
args : List Arg
/-- Instance implicit arguments collected so far. -/
instMVars : Array MVarId := #[]
/-- Position of the next argument to be processed. We use it to decide whether the argument is the motive or a discriminant. -/
idx : Nat := 0
/-- Store the metavariable used to represent the motive that will be computed at `finalize`. -/
motive? : Option Expr := none
abbrev M := ReaderT Context $ StateRefT State TermElabM
/-- Infer the `motive` using the expected type by `kabstract`ing the discriminants. -/
def mkMotive (discrs : Array Expr) (expectedType : Expr) : MetaM Expr := do
discrs.foldrM (init := expectedType) fun discr motive => do
let discr ← instantiateMVars discr
let motiveBody ← kabstract motive discr
/- We use `transform (usedLetOnly := true)` to eliminate unnecessary let-expressions. -/
let discrType ← transform (usedLetOnly := true) (← instantiateMVars (← inferType discr))
return Lean.mkLambda (← mkFreshBinderName) BinderInfo.default discrType motiveBody
/--
If the eliminator is over-applied, we "revert" the extra arguments.
Returns the function with the reverted arguments applied and the new generalized expected type.
-/
def revertArgs (args : List Arg) (f : Expr) (expectedType : Expr) : TermElabM (Expr × Expr) := do
let (xs, expectedType) ← args.foldrM (init := ([], expectedType)) fun arg (xs, expectedType) => do
let val ←
match arg with
| .expr val => pure val
| .stx stx => elabTerm stx none
let val ← instantiateMVars val
let expectedTypeBody ← kabstract expectedType val
/- We use `transform (usedLetOnly := true)` to eliminate unnecessary let-expressions. -/
let valType ← transform (usedLetOnly := true) (← instantiateMVars (← inferType val))
return (val :: xs, mkForall (← mkFreshBinderName) BinderInfo.default valType expectedTypeBody)
return (xs.foldl .app f, expectedType)
/--
Construct the resulting application after all discriminants have been elaborated, and we have
consumed as many given arguments as possible.
-/
def finalize : M Expr := do
unless (← get).namedArgs.isEmpty do
throwError "failed to elaborate eliminator, unused named arguments: {(← get).namedArgs.map (·.name)}"
let some motive := (← get).motive?
| throwError "failed to elaborate eliminator, insufficient number of arguments"
trace[Elab.app.elab_as_elim] "motive: {motive}"
forallTelescope (← get).fType fun xs fType => do
trace[Elab.app.elab_as_elim] "xs: {xs}"
let mut expectedType := (← read).expectedType
trace[Elab.app.elab_as_elim] "expectedType:{indentD expectedType}"
let throwInsufficient := do
throwError "failed to elaborate eliminator, insufficient number of arguments, expected type:{indentExpr expectedType}"
let mut f := (← get).f
if xs.size > 0 then
-- under-application, specialize the expected type using `xs`
-- Note: if we ever wanted to support optParams and autoParams, this is where it could be.
assert! (← get).args.isEmpty
for x in xs do
let .forallE _ t b _ ← whnf expectedType | throwInsufficient
unless ← fullApproxDefEq <| isDefEq t (← inferType x) do
-- We can't assume that these binding domains were supposed to line up, so report insufficient arguments
throwInsufficient
expectedType := b.instantiate1 x
trace[Elab.app.elab_as_elim] "xs after specialization of expected type: {xs}"
else
-- over-application, simulate `revert` while generalizing the values of these arguments in the expected type
(f, expectedType) ← revertArgs (← get).args f expectedType
unless ← isTypeCorrect expectedType do
throwError "failed to elaborate eliminator, after generalizing over-applied arguments, expected type is type incorrect:{indentExpr expectedType}"
trace[Elab.app.elab_as_elim] "expectedType after processing:{indentD expectedType}"
let result := mkAppN f xs
trace[Elab.app.elab_as_elim] "result:{indentD result}"
unless fType.getAppFn == (← get).motive? do
throwError "Internal error, eliminator target type isn't an application of the motive"
let discrs := fType.getAppArgs
trace[Elab.app.elab_as_elim] "discrs: {discrs}"
let motiveVal ← mkMotive discrs expectedType
unless (← isTypeCorrect motiveVal) do
throwError "failed to elaborate eliminator, motive is not type correct:{indentD motiveVal}"
unless (← isDefEq motive motiveVal) do
throwError "failed to elaborate eliminator, invalid motive{indentExpr motiveVal}"
synthesizeAppInstMVars (← get).instMVars result
trace[Elab.app.elab_as_elim] "completed motive:{indentD motive}"
let result ← mkLambdaFVars xs (← instantiateMVars result)
return result
/--
Return the next argument to be processed.
The result is `.none` if it is an implicit argument which was not provided using a named argument.
The result is `.undef` if `args` is empty and `namedArgs` does contain an entry for `binderName`.
-/
def getNextArg? (binderName : Name) (binderInfo : BinderInfo) : M (LOption Arg) := do
match findBinderName? (← get).namedArgs binderName with
| some namedArg =>
modify fun s => { s with namedArgs := eraseNamedArg s.namedArgs binderName }
return .some namedArg.val
| none =>
if binderInfo.isExplicit then
match (← get).args with
| [] => return .undef
| arg :: args =>
modify fun s => { s with args }
return .some arg
else
return .none
/-- Set the `motive` field in the state. -/
def setMotive (motive : Expr) : M Unit :=
modify fun s => { s with motive? := motive }
/-- Elaborate the given argument with the given expected type. -/
private def elabArg (arg : Arg) (argExpectedType : Expr) : M Expr := do
match arg with
| Arg.expr val => ensureArgType (← get).f val argExpectedType
| Arg.stx stx =>
let val ← elabTerm stx argExpectedType
withRef stx <| ensureArgType (← get).f val argExpectedType
/-- Save information for producing error messages. -/
def saveArgInfo (arg : Expr) (binderName : Name) : M Unit := do
if arg.isMVar then
registerMVarArgName arg.mvarId! binderName
/-- Create an implicit argument using the given `BinderInfo`. -/
def mkImplicitArg (argExpectedType : Expr) (bi : BinderInfo) : M Expr := do
let arg ← mkFreshExprMVar argExpectedType (if bi.isInstImplicit then .synthetic else .natural)
if bi.isInstImplicit then
modify fun s => { s with instMVars := s.instMVars.push arg.mvarId! }
return arg