-
Notifications
You must be signed in to change notification settings - Fork 434
/
ExprDefEq.lean
2212 lines (2008 loc) · 93.2 KB
/
ExprDefEq.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.Meta.Offset
import Lean.Meta.UnificationHint
import Lean.Util.OccursCheck
namespace Lean.Meta
register_builtin_option backward.isDefEq.lazyProjDelta : Bool := {
defValue := true
group := "backward compatibility"
descr := "use lazy delta reduction when solving unification constrains of the form `(f a).i =?= (g b).i`"
}
register_builtin_option backward.isDefEq.lazyWhnfCore : Bool := {
defValue := true
group := "backward compatibility"
descr := "specifies transparency mode when normalizing constraints of the form `(f a).i =?= s`, if `true` only reducible definitions and instances are unfolded when reducing `f a`. Otherwise, the default setting is used"
}
/--
Return `true` if `e` is of the form `fun (x_1 ... x_n) => ?m y_1 ... y_k)`, and `?m` is unassigned.
Remark: `n`, `k` may be 0.
This function is used to filter unification problems in
`isDefEqArgs`/`isDefEqEtaStruct` where we can assign proofs.
If one side is of the form described above, then we can likely assign `?m`.
But it it's not, we would most likely apply proof irrelevance, which is
usually very expensive since it needs to unify the types as well.
-/
def isAbstractedUnassignedMVar : Expr → MetaM Bool
| .lam _ _ b _ => isAbstractedUnassignedMVar b
| .app a _ => isAbstractedUnassignedMVar a
| .mvar mvarId => do
if (← mvarId.isReadOnlyOrSyntheticOpaque) then
pure false
else if (← mvarId.isAssigned) then
pure false
else
pure true
| _ => pure false
/--
Return true if `b` is of the form `mk a.1 ... a.n`, and `a` is not a constructor application.
If `a` and `b` are constructor applications, the method returns `false` to force `isDefEq` to use `isDefEqArgs`.
For example, suppose we are trying to solve the constraint
```
Fin.mk ?n ?h =?= Fin.mk n h
```
If this method is applied, the constraints are reduced to
```
n =?= (Fin.mk ?n ?h).1
h =?= (Fin.mk ?n ?h).2
```
The first constraint produces the assignment `?n := n`. Then, the second constraint is solved using proof irrelevance without
assigning `?h`.
TODO: investigate better solutions for the proof irrelevance issue. The problem above can happen is other scenarios.
That is, proof irrelevance may prevent us from performing desired mvar assignments.
-/
private def isDefEqEtaStruct (a b : Expr) : MetaM Bool := do
matchConstCtor b.getAppFn (fun _ => return false) fun ctorVal us => do
if (← useEtaStruct ctorVal.induct) then
matchConstCtor a.getAppFn (fun _ => go ctorVal us) fun _ _ => return false
else
return false
where
go ctorVal us := do
if ctorVal.numParams + ctorVal.numFields != b.getAppNumArgs then
trace[Meta.isDefEq.eta.struct] "failed, insufficient number of arguments at{indentExpr b}"
return false
else
if !isStructureLike (← getEnv) ctorVal.induct then
trace[Meta.isDefEq.eta.struct] "failed, type is not a structure{indentExpr b}"
return false
else if (← isDefEq (← inferType a) (← inferType b)) then
checkpointDefEq do
let args := b.getAppArgs
let params := args[:ctorVal.numParams].toArray
for i in [ctorVal.numParams : args.size] do
let j := i - ctorVal.numParams
let proj ← mkProjFn ctorVal us params j a
if ← isProof proj then
unless ← isAbstractedUnassignedMVar args[i]! do
-- Skip expensive unification problem that is likely solved
-- using proof irrelevance. We already know that `proj` and
-- `args[i]!` have the same type, so they're defeq in any case.
-- See comment at `isAbstractedUnassignedMVar`.
continue
trace[Meta.isDefEq.eta.struct] "{a} =?= {b} @ [{j}], {proj} =?= {args[i]!}"
unless (← isDefEq proj args[i]!) do
trace[Meta.isDefEq.eta.struct] "failed, unexpect arg #{i}, projection{indentExpr proj}\nis not defeq to{indentExpr args[i]!}"
return false
return true
else
return false
/--
Try to solve `a := (fun x => t) =?= b` by eta-expanding `b`,
resulting in `t =?= b x` (with a fresh free variable `x`).
Remark: eta-reduction is not a good alternative even in a system without universe cumulativity like Lean.
Example:
```
(fun x : A => f ?m) =?= f
```
The left-hand side of the constraint above it not eta-reduced because `?m` is a metavariable.
Note: we do not backtrack after applying η-expansion anymore.
There is no case where `(fun x => t) =?= b` unifies, but `t =?= b x` does not.
Backtracking after η-expansion results in lots of duplicate δ-reductions,
because we can δ-reduce `a` before and after the η-expansion.
The fresh free variable `x` also busts the cache.
See https://github.com/leanprover/lean4/pull/2002 -/
private def isDefEqEta (a b : Expr) : MetaM LBool := do
if a.isLambda && !b.isLambda then
let bType ← inferType b
let bType ← whnfD bType
match bType with
| .forallE n d _ c =>
let b' := mkLambda n c d (mkApp b (mkBVar 0))
toLBoolM <| Meta.isExprDefEqAux a b'
| _ => return .undef
else
return .undef
/-- Support for `Lean.reduceBool` and `Lean.reduceNat` -/
def isDefEqNative (s t : Expr) : MetaM LBool := do
let isDefEq (s t) : MetaM LBool := toLBoolM <| Meta.isExprDefEqAux s t
let s? ← reduceNative? s
let t? ← reduceNative? t
match s?, t? with
| some s, some t => isDefEq s t
| some s, none => isDefEq s t
| none, some t => isDefEq s t
| none, none => pure LBool.undef
/-- Support for reducing Nat basic operations. -/
def isDefEqNat (s t : Expr) : MetaM LBool := do
let isDefEq (s t) : MetaM LBool := toLBoolM <| Meta.isExprDefEqAux s t
if s.hasFVar || s.hasMVar || t.hasFVar || t.hasMVar then
pure LBool.undef
else
let s? ← reduceNat? s
let t? ← reduceNat? t
match s?, t? with
| some s, some t => isDefEq s t
| some s, none => isDefEq s t
| none, some t => isDefEq s t
| none, none => pure LBool.undef
/-- Support for constraints of the form `("..." =?= String.mk cs)` -/
def isDefEqStringLit (s t : Expr) : MetaM LBool := do
let isDefEq (s t) : MetaM LBool := toLBoolM <| Meta.isExprDefEqAux s t
if s.isStringLit && t.isAppOf ``String.mk then
isDefEq s.toCtorIfLit t
else if s.isAppOf `String.mk && t.isStringLit then
isDefEq s t.toCtorIfLit
else
pure LBool.undef
/--
Return `true` if `e` is of the form `fun (x_1 ... x_n) => ?m x_1 ... x_n)`, and `?m` is unassigned.
Remark: `n` may be 0. -/
def isEtaUnassignedMVar (e : Expr) : MetaM Bool := do
match e.etaExpanded? with
| some (Expr.mvar mvarId) =>
if (← mvarId.isReadOnlyOrSyntheticOpaque) then
pure false
else if (← mvarId.isAssigned) then
pure false
else
pure true
| _ => pure false
private def trySynthPending (e : Expr) : MetaM Bool := do
let mvarId? ← getStuckMVar? e
match mvarId? with
| some mvarId => Meta.synthPending mvarId
| none => pure false
/--
Result type for `isDefEqArgsFirstPass`.
-/
inductive DefEqArgsFirstPassResult where
/--
Failed to establish that explicit arguments are def-eq.
Remark: higher output parameters, and parameters that depend on them
are postponed.
-/
| failed
/--
Succeeded. The array `postponedImplicit` contains the position
of the implicit arguments for which def-eq has been postponed.
`postponedHO` contains the higher order output parameters, and parameters
that depend on them. They should be processed after the implicit ones.
`postponedHO` is used to handle applications involving functions that
contain higher order output parameters. Example:
```lean
getElem :
{cont : Type u_1} → {idx : Type u_2} → {elem : Type u_3} →
{dom : cont → idx → Prop} → [self : GetElem cont idx elem dom] →
(xs : cont) → (i : idx) → (h : dom xs i) → elem
```
The arguments `dom` and `h` must be processed after all implicit arguments
otherwise higher-order unification problems are generated. See issue #1299,
when trying to solve
```
getElem ?a ?i ?h =?= getElem a i (Fin.val_lt_of_le i ...)
```
we have to solve the constraint
```
?dom a i.val =?= LT.lt i.val (Array.size a)
```
by solving after the instance has been synthesized, we reduce this constraint to
a simple check.
-/
| ok (postponedImplicit : Array Nat) (postponedHO : Array Nat)
/--
First pass for `isDefEqArgs`. We unify explicit arguments, *and* easy cases
Here, we say a case is easy if it is of the form
?m =?= t
or
t =?= ?m
where `?m` is unassigned.
These easy cases are not just an optimization. When
`?m` is a function, by assigning it to t, we make sure
a unification constraint (in the explicit part)
```
?m t =?= f s
```
is not higher-order.
We also handle the eta-expanded cases:
```
fun x₁ ... xₙ => ?m x₁ ... xₙ =?= t
t =?= fun x₁ ... xₙ => ?m x₁ ... xₙ
```
This is important because type inference often produces
eta-expanded terms, and without this extra case, we could
introduce counter intuitive behavior.
Pre: `paramInfo.size <= args₁.size = args₂.size`
See `DefEqArgsFirstPassResult` for additional information.
-/
private def isDefEqArgsFirstPass
(paramInfo : Array ParamInfo) (args₁ args₂ : Array Expr) : MetaM DefEqArgsFirstPassResult := do
let mut postponedImplicit := #[]
let mut postponedHO := #[]
for i in [:paramInfo.size] do
let info := paramInfo[i]!
let a₁ := args₁[i]!
let a₂ := args₂[i]!
if info.dependsOnHigherOrderOutParam || info.higherOrderOutParam then
trace[Meta.isDefEq] "found messy {a₁} =?= {a₂}"
postponedHO := postponedHO.push i
else if info.isExplicit then
if info.isProp then
unless ← isAbstractedUnassignedMVar a₁ <||> isAbstractedUnassignedMVar a₂ do
-- Skip expensive unification problem that is likely solved
-- using proof irrelevance. We already know that `a₁` and
-- `a₂` have the same type, so they're defeq in any case.
-- See comment at `isAbstractedUnassignedMVar`.
continue
unless (← Meta.isExprDefEqAux a₁ a₂) do
return .failed
else if (← isEtaUnassignedMVar a₁ <||> isEtaUnassignedMVar a₂) then
unless (← Meta.isExprDefEqAux a₁ a₂) do
return .failed
else
if info.isProp then
unless ← isAbstractedUnassignedMVar a₁ <||> isAbstractedUnassignedMVar a₂ do
-- Skip expensive unification problem that is likely solved
-- using proof irrelevance. We already know that `a₁` and
-- `a₂` have the same type, so they're defeq in any case.
-- See comment at `isAbstractedUnassignedMVar`.
continue
postponedImplicit := postponedImplicit.push i
return .ok postponedImplicit postponedHO
private partial def isDefEqArgs (f : Expr) (args₁ args₂ : Array Expr) : MetaM Bool := do
unless args₁.size == args₂.size do return false
let finfo ← getFunInfoNArgs f args₁.size
let .ok postponedImplicit postponedHO ← isDefEqArgsFirstPass finfo.paramInfo args₁ args₂ | pure false
-- finfo.paramInfo.size may be smaller than args₁.size
for i in [finfo.paramInfo.size:args₁.size] do
unless (← Meta.isExprDefEqAux args₁[i]! args₂[i]!) do
return false
for i in postponedImplicit do
/- Second pass: unify implicit arguments.
In the second pass, we make sure we are unfolding at
least non reducible definitions (default setting). -/
let a₁ := args₁[i]!
let a₂ := args₂[i]!
let info := finfo.paramInfo[i]!
if info.isInstImplicit then
discard <| trySynthPending a₁
discard <| trySynthPending a₂
unless (← withAtLeastTransparency TransparencyMode.default <| Meta.isExprDefEqAux a₁ a₂) do
return false
for i in postponedHO do
let a₁ := args₁[i]!
let a₂ := args₂[i]!
let info := finfo.paramInfo[i]!
if info.isInstImplicit then
unless (← withAtLeastTransparency TransparencyMode.default <| Meta.isExprDefEqAux a₁ a₂) do
return false
else
unless (← Meta.isExprDefEqAux a₁ a₂) do
return false
return true
/--
Check whether the types of the free variables at `fvars` are
definitionally equal to the types at `ds₂`.
Pre: `fvars.size == ds₂.size`
This method also updates the set of local instances, and invokes
the continuation `k` with the updated set.
We can't use `withNewLocalInstances` because the `isDeq fvarType d₂`
may use local instances. -/
@[specialize] partial def isDefEqBindingDomain (fvars : Array Expr) (ds₂ : Array Expr) (k : MetaM Bool) : MetaM Bool :=
let rec loop (i : Nat) := do
if h : i < fvars.size then do
let fvar := fvars.get ⟨i, h⟩
let fvarDecl ← getFVarLocalDecl fvar
let fvarType := fvarDecl.type
let d₂ := ds₂[i]!
if (← Meta.isExprDefEqAux fvarType d₂) then
match (← isClass? fvarType) with
| some className => withNewLocalInstance className fvar <| loop (i+1)
| none => loop (i+1)
else
pure false
else
k
loop 0
/--
Auxiliary function for `isDefEqBinding` for handling binders `forall/fun`.
It accumulates the new free variables in `fvars`, and declare them at `lctx`.
We use the domain types of `e₁` to create the new free variables.
We store the domain types of `e₂` at `ds₂`.
-/
private partial def isDefEqBindingAux (lctx : LocalContext) (fvars : Array Expr) (e₁ e₂ : Expr) (ds₂ : Array Expr) : MetaM Bool :=
let process (n : Name) (d₁ d₂ b₁ b₂ : Expr) : MetaM Bool := do
let d₁ := d₁.instantiateRev fvars
let d₂ := d₂.instantiateRev fvars
let fvarId ← mkFreshFVarId
let lctx := lctx.mkLocalDecl fvarId n d₁
let fvars := fvars.push (mkFVar fvarId)
isDefEqBindingAux lctx fvars b₁ b₂ (ds₂.push d₂)
match e₁, e₂ with
| Expr.forallE n d₁ b₁ _, Expr.forallE _ d₂ b₂ _ => process n d₁ d₂ b₁ b₂
| Expr.lam n d₁ b₁ _, Expr.lam _ d₂ b₂ _ => process n d₁ d₂ b₁ b₂
| _, _ =>
withReader (fun ctx => { ctx with lctx := lctx }) do
isDefEqBindingDomain fvars ds₂ do
Meta.isExprDefEqAux (e₁.instantiateRev fvars) (e₂.instantiateRev fvars)
@[inline] private def isDefEqBinding (a b : Expr) : MetaM Bool := do
let lctx ← getLCtx
isDefEqBindingAux lctx #[] a b #[]
private def checkTypesAndAssign (mvar : Expr) (v : Expr) : MetaM Bool :=
withTraceNodeBefore `Meta.isDefEq.assign.checkTypes (return m!"({mvar} : {← inferType mvar}) := ({v} : {← inferType v})") do
if !mvar.isMVar then
trace[Meta.isDefEq.assign.checkTypes] "metavariable expected"
return false
else
-- must check whether types are definitionally equal or not, before assigning and returning true
let mvarType ← inferType mvar
let vType ← inferType v
if (← withTransparency TransparencyMode.default <| Meta.isExprDefEqAux mvarType vType) then
mvar.mvarId!.assign v
pure true
else
pure false
/--
Auxiliary method for solving constraints of the form `?m xs := v`.
It creates a lambda using `mkLambdaFVars ys v`, where `ys` is a superset of `xs`.
`ys` is often equal to `xs`. It is a bigger when there are let-declaration dependencies in `xs`.
For example, suppose we have `xs` of the form `#[a, c]` where
```
a : Nat
b : Nat := f a
c : b = a
```
In this scenario, the type of `?m` is `(x1 : Nat) -> (x2 : f x1 = x1) -> C[x1, x2]`,
and type of `v` is `C[a, c]`. Note that, `?m a c` is type correct since `f a = a` is definitionally equal
to the type of `c : b = a`, and the type of `?m a c` is equal to the type of `v`.
Note that `fun xs => v` is the term `fun (x1 : Nat) (x2 : b = x1) => v` which has type
`(x1 : Nat) -> (x2 : b = x1) -> C[x1, x2]` which is not definitionally equal to the type of `?m`,
and may not even be type correct.
The issue here is that we are not capturing the `let`-declarations.
This method collects let-declarations `y` occurring between `xs[0]` and `xs.back` s.t.
some `x` in `xs` depends on `y`.
`ys` is the `xs` with these extra let-declarations included.
In the example above, `ys` is `#[a, b, c]`, and `mkLambdaFVars ys v` produces
`fun a => let b := f a; fun (c : b = a) => v` which has a type definitionally equal to the type of `?m`.
Recall that the method `checkAssignment` ensures `v` does not contain offending `let`-declarations.
This method assumes that for any `xs[i]` and `xs[j]` where `i < j`, we have that `index of xs[i]` < `index of xs[j]`.
where the index is the position in the local context.
-/
private partial def mkLambdaFVarsWithLetDeps (xs : Array Expr) (v : Expr) : MetaM (Option Expr) := do
if not (← hasLetDeclsInBetween) then
mkLambdaFVars xs v (etaReduce := true)
else
let ys ← addLetDeps
mkLambdaFVars ys v (etaReduce := true)
where
/-- Return true if there are let-declarions between `xs[0]` and `xs[xs.size-1]`.
We use it a quick-check to avoid the more expensive collection procedure. -/
hasLetDeclsInBetween : MetaM Bool := do
let check (lctx : LocalContext) : Bool := Id.run do
let start := lctx.getFVar! xs[0]! |>.index
let stop := lctx.getFVar! xs.back |>.index
for i in [start+1:stop] do
match lctx.getAt? i with
| some localDecl =>
if localDecl.isLet then
return true
| _ => pure ()
return false
if xs.size <= 1 then
return false
else
return check (← getLCtx)
/-- Traverse `e` and stores in the state `NameHashSet` any let-declaration with index greater than `(← read)`.
The context `Nat` is the position of `xs[0]` in the local context. -/
collectLetDeclsFrom (e : Expr) : ReaderT Nat (StateRefT FVarIdHashSet MetaM) Unit := do
let rec visit (e : Expr) : MonadCacheT Expr Unit (ReaderT Nat (StateRefT FVarIdHashSet MetaM)) Unit :=
checkCache e fun _ => do
match e with
| .forallE _ d b _ => visit d; visit b
| .lam _ d b _ => visit d; visit b
| .letE _ t v b _ => visit t; visit v; visit b
| .app f a => visit f; visit a
| .mdata _ b => visit b
| .proj _ _ b => visit b
| .fvar fvarId =>
let localDecl ← fvarId.getDecl
if localDecl.isLet && localDecl.index > (← read) then
modify fun s => s.insert localDecl.fvarId
| _ => pure ()
visit (← instantiateMVars e) |>.run
/--
Auxiliary definition for traversing all declarations between `xs[0]` ... `xs.back` backwards.
The `Nat` argument is the current position in the local context being visited, and it is less than
or equal to the position of `xs.back` in the local context.
The `Nat` context `(← read)` is the position of `xs[0]` in the local context.
-/
collectLetDepsAux : Nat → ReaderT Nat (StateRefT FVarIdHashSet MetaM) Unit
| 0 => return ()
| i+1 => do
if i+1 == (← read) then
return ()
else
match (← getLCtx).getAt? (i+1) with
| none => collectLetDepsAux i
| some localDecl =>
if (← get).contains localDecl.fvarId then
collectLetDeclsFrom localDecl.type
match localDecl.value? with
| some val => collectLetDeclsFrom val
| _ => pure ()
collectLetDepsAux i
/-- Computes the set `ys`. It is a set of `FVarId`s, -/
collectLetDeps : MetaM FVarIdHashSet := do
let lctx ← getLCtx
let start := lctx.getFVar! xs[0]! |>.index
let stop := lctx.getFVar! xs.back |>.index
let s := xs.foldl (init := {}) fun s x => s.insert x.fvarId!
let (_, s) ← collectLetDepsAux stop |>.run start |>.run s
return s
/-- Computes the array `ys` containing let-decls between `xs[0]` and `xs.back` that
some `x` in `xs` depends on. -/
addLetDeps : MetaM (Array Expr) := do
let lctx ← getLCtx
let s ← collectLetDeps
/- Convert `s` into the array `ys` -/
let start := lctx.getFVar! xs[0]! |>.index
let stop := lctx.getFVar! xs.back |>.index
let mut ys := #[]
for i in [start:stop+1] do
match lctx.getAt? i with
| none => pure ()
| some localDecl =>
if s.contains localDecl.fvarId then
ys := ys.push localDecl.toExpr
return ys
/-!
Each metavariable is declared in a particular local context.
We use the notation `C |- ?m : t` to denote a metavariable `?m` that
was declared at the local context `C` with type `t` (see `MetavarDecl`).
We also use `?m@C` as a shorthand for `C |- ?m : t` where `t` is the type of `?m`.
The following method process the unification constraint
?m@C a₁ ... aₙ =?= t
We say the unification constraint is a pattern IFF
1) `a₁ ... aₙ` are pairwise distinct free variables that are *not* let-variables.
2) `a₁ ... aₙ` are not in `C`
3) `t` only contains free variables in `C` and/or `{a₁, ..., aₙ}`
4) For every metavariable `?m'@C'` occurring in `t`, `C'` is a subprefix of `C`
5) `?m` does not occur in `t`
Claim: we don't have to check free variable declarations. That is,
if `t` contains a reference to `x : A := v`, we don't need to check `v`.
Reason: The reference to `x` is a free variable, and it must be in `C` (by 1 and 3).
If `x` is in `C`, then any metavariable occurring in `v` must have been defined in a strict subprefix of `C`.
So, condition 4 and 5 are satisfied.
If the conditions above have been satisfied, then the
solution for the unification constrain is
?m := fun a₁ ... aₙ => t
Now, we consider some workarounds/approximations.
A1) Suppose `t` contains a reference to `x : A := v` and `x` is not in `C` (failed condition 3)
(precise) solution: unfold `x` in `t`.
A2) Suppose some `aᵢ` is in `C` (failed condition 2)
(approximated) solution (when `config.quasiPatternApprox` is set to true) :
ignore condition and also use
?m := fun a₁ ... aₙ => t
Here is an example where this approximation fails:
Given `C` containing `a : nat`, consider the following two constraints
?m@C a =?= a
?m@C b =?= a
If we use the approximation in the first constraint, we get
?m := fun x => x
when we apply this solution to the second one we get a failure.
IMPORTANT: When applying this approximation we need to make sure the
abstracted term `fun a₁ ... aₙ => t` is type correct. The check
can only be skipped in the pattern case described above. Consider
the following example. Given the local context
(α : Type) (a : α)
we try to solve
?m α =?= @id α a
If we use the approximation above we obtain:
?m := (fun α' => @id α' a)
which is a type incorrect term. `a` has type `α` but it is expected to have
type `α'`.
The problem occurs because the right hand side contains a free variable
`a` that depends on the free variable `α` being abstracted. Note that
this dependency cannot occur in patterns.
We can address this by type checking
the term after abstraction. This is not a significant performance
bottleneck because this case doesn't happen very often in practice
(262 times when compiling stdlib on Jan 2018). The second example
is trickier, but it also occurs less frequently (8 times when compiling
stdlib on Jan 2018, and all occurrences were at Init/Control when
we define monads and auxiliary combinators for them).
We considered three options for the addressing the issue on the second example:
A3) `a₁ ... aₙ` are not pairwise distinct (failed condition 1).
In Lean3, we would try to approximate this case using an approach similar to A2.
However, this approximation complicates the code, and is never used in the
Lean3 stdlib and mathlib.
A4) `t` contains a metavariable `?m'@C'` where `C'` is not a subprefix of `C`.
If `?m'` is assigned, we substitute.
If not, we create an auxiliary metavariable with a smaller scope.
Actually, we let `elimMVarDeps` at `MetavarContext.lean` to perform this step.
A5) If some `aᵢ` is not a free variable,
then we use first-order unification (if `config.foApprox` is set to true)
?m a_1 ... a_i a_{i+1} ... a_{i+k} =?= f b_1 ... b_k
reduces to
?M a_1 ... a_i =?= f
a_{i+1} =?= b_1
...
a_{i+k} =?= b_k
A6) If (m =?= v) is of the form
?m a_1 ... a_n =?= ?m b_1 ... b_k
then we use first-order unification (if `config.foApprox` is set to true)
A7) When `foApprox`, we may use another approximation (`constApprox`) for solving constraints of the form
```
?m s₁ ... sₙ =?= t
```
where `s₁ ... sₙ` are arbitrary terms. We solve them by assigning the constant function to `?m`.
```
?m := fun _ ... _ => t
```
In general, this approximation may produce bad solutions, and may prevent coercions from being tried.
For example, consider the term `pure (x > 0)` with inferred type `?m Prop` and expected type `IO Bool`.
In this situation, the
elaborator generates the unification constraint
```
?m Prop =?= IO Bool
```
It is not a higher-order pattern, nor first-order approximation is applicable. However, constant approximation
produces the bogus solution `?m := fun _ => IO Bool`, and prevents the system from using the coercion from
the decidable proposition `x > 0` to `Bool`.
On the other hand, the constant approximation is desirable for elaborating the term
```
let f (x : _) := pure "hello"; f ()
```
with expected type `IO String`.
In this example, the following unification constraint is generated.
```
?m () String =?= IO String
```
It is not a higher-order pattern, first-order approximation reduces it to
```
?m () =?= IO
```
which fails to be solved. However, constant approximation solves it by assigning
```
?m := fun _ => IO
```
Note that `f`s type is `(x : ?α) -> ?m x String`. The metavariable `?m` may depend on `x`.
If `constApprox` is set to true, we use constant approximation. Otherwise, we use a heuristic to decide
whether we should apply it or not. The heuristic is based on observing where the constraints above come from.
In the first example, the constraint `?m Prop =?= IO Bool` come from polymorphic method where `?m` is expected to
be a **function** of type `Type -> Type`. In the second example, the first argument of `?m` is used to model
a **potential** dependency on `x`. By using constant approximation here, we are just saying the type of `f`
does **not** depend on `x`. We claim this is a reasonable approximation in practice. Moreover, it is expected
by any functional programmer used to non-dependently type languages (e.g., Haskell).
We distinguish the two cases above by using the field `numScopeArgs` at `MetavarDecl`. This field tracks
how many metavariable arguments are representing dependencies.
-/
def mkAuxMVar (lctx : LocalContext) (localInsts : LocalInstances) (type : Expr) (numScopeArgs : Nat := 0) : MetaM Expr := do
mkFreshExprMVarAt lctx localInsts type MetavarKind.natural Name.anonymous numScopeArgs
namespace CheckAssignment
builtin_initialize checkAssignmentExceptionId : InternalExceptionId ← registerInternalExceptionId `checkAssignment
builtin_initialize outOfScopeExceptionId : InternalExceptionId ← registerInternalExceptionId `outOfScope
structure State where
cache : ExprStructMap Expr := {}
structure Context where
mvarId : MVarId
mvarDecl : MetavarDecl
fvars : Array Expr
hasCtxLocals : Bool
rhs : Expr
abbrev CheckAssignmentM := ReaderT Context $ StateRefT State MetaM
def throwCheckAssignmentFailure : CheckAssignmentM α :=
throw <| Exception.internal checkAssignmentExceptionId
def throwOutOfScopeFVar : CheckAssignmentM α :=
throw <| Exception.internal outOfScopeExceptionId
private def findCached? (e : Expr) : CheckAssignmentM (Option Expr) := do
return (← get).cache.find? e
private def cache (e r : Expr) : CheckAssignmentM Unit := do
modify fun s => { s with cache := s.cache.insert e r }
instance : MonadCache Expr Expr CheckAssignmentM where
findCached? := findCached?
cache := cache
private def addAssignmentInfo (msg : MessageData) : CheckAssignmentM MessageData := do
let ctx ← read
return m!"{msg} @ {mkMVar ctx.mvarId} {ctx.fvars} := {ctx.rhs}"
@[inline] def run (x : CheckAssignmentM Expr) (mvarId : MVarId) (fvars : Array Expr) (hasCtxLocals : Bool) (v : Expr) : MetaM (Option Expr) := do
let mvarDecl ← mvarId.getDecl
let ctx := { mvarId := mvarId, mvarDecl := mvarDecl, fvars := fvars, hasCtxLocals := hasCtxLocals, rhs := v : Context }
let x : CheckAssignmentM (Option Expr) :=
catchInternalIds [outOfScopeExceptionId, checkAssignmentExceptionId]
(do let e ← x; return some e)
(fun _ => pure none)
x.run ctx |>.run' {}
mutual
partial def checkFVar (fvar : Expr) : CheckAssignmentM Expr := do
let ctxMeta ← readThe Meta.Context
let ctx ← read
if ctx.mvarDecl.lctx.containsFVar fvar then
pure fvar
else
let lctx := ctxMeta.lctx
match lctx.findFVar? fvar with
| some (.ldecl (value := v) ..) => check v
| _ =>
if ctx.fvars.contains fvar then pure fvar
else
traceM `Meta.isDefEq.assign.outOfScopeFVar do addAssignmentInfo fvar
throwOutOfScopeFVar
partial def checkMVar (mvar : Expr) : CheckAssignmentM Expr := do
let mvarId := mvar.mvarId!
let ctx ← read
if mvarId == ctx.mvarId then
traceM `Meta.isDefEq.assign.occursCheck <| addAssignmentInfo "occurs check failed"
throwCheckAssignmentFailure
if let some v ← getExprMVarAssignment? mvarId then
return (← check v)
let some mvarDecl ← mvarId.findDecl?
| throwUnknownMVar mvarId
if ctx.hasCtxLocals then
throwCheckAssignmentFailure -- It is not a pattern, then we fail and fall back to FO unification
if let some d ← getDelayedMVarAssignment? mvarId then
-- we must perform occurs-check at `d.mvarIdPending`
unless (← occursCheck ctx.mvarId (mkMVar d.mvarIdPending)) do
traceM `Meta.isDefEq.assign.occursCheck <| addAssignmentInfo "occurs check failed"
throwCheckAssignmentFailure
if mvarDecl.lctx.isSubPrefixOf ctx.mvarDecl.lctx ctx.fvars then
/- The local context of `mvar` - free variables being abstracted is a subprefix of the metavariable being assigned.
We "subtract" variables being abstracted because we use `elimMVarDeps` -/
return mvar
if mvarDecl.depth != (← getMCtx).depth || mvarDecl.kind.isSyntheticOpaque then
traceM `Meta.isDefEq.assign.readOnlyMVarWithBiggerLCtx <| addAssignmentInfo (mkMVar mvarId)
throwCheckAssignmentFailure
let ctxMeta ← readThe Meta.Context
unless ctxMeta.config.ctxApprox && ctx.mvarDecl.lctx.isSubPrefixOf mvarDecl.lctx do
traceM `Meta.isDefEq.assign.readOnlyMVarWithBiggerLCtx <| addAssignmentInfo (mkMVar mvarId)
throwCheckAssignmentFailure
/- Create an auxiliary metavariable with a smaller context and "checked" type.
Note that `mvarType` may be different from `mvarDecl.type`. Example: `mvarType` contains
a metavariable that we also need to reduce the context.
We remove from `ctx.mvarDecl.lctx` any variable that is not in `mvarDecl.lctx`
or in `ctx.fvars`. We don't need to remove the ones in `ctx.fvars` because
`elimMVarDeps` will take care of them.
First, we collect `toErase` the variables that need to be erased.
Notat that if a variable is `ctx.fvars`, but it depends on variable at `toErase`,
we must also erase it.
-/
let toErase ← mvarDecl.lctx.foldlM (init := #[]) fun toErase localDecl => do
if ctx.mvarDecl.lctx.contains localDecl.fvarId then
return toErase
else if ctx.fvars.any fun fvar => fvar.fvarId! == localDecl.fvarId then
if (← findLocalDeclDependsOn localDecl fun fvarId => toErase.contains fvarId) then
-- localDecl depends on a variable that will be erased. So, we must add it to `toErase` too
return toErase.push localDecl.fvarId
else
return toErase
else
return toErase.push localDecl.fvarId
let lctx := toErase.foldl (init := mvarDecl.lctx) fun lctx toEraseFVar =>
lctx.erase toEraseFVar
/- Compute new set of local instances. -/
let localInsts := mvarDecl.localInstances.filter fun localInst => !toErase.contains localInst.fvar.fvarId!
let mvarType ← check mvarDecl.type
let newMVar ← mkAuxMVar lctx localInsts mvarType mvarDecl.numScopeArgs
mvarId.assign newMVar
return newMVar
/--
Auxiliary function used to "fix" subterms of the form `?m x_1 ... x_n` where `x_i`s are free variables,
and one of them is out-of-scope.
See `Expr.app` case at `check`.
If `ctxApprox` is true, then we solve this case by creating a fresh metavariable ?n with the correct scope,
an assigning `?m := fun _ ... _ => ?n` -/
partial def assignToConstFun (mvar : Expr) (numArgs : Nat) (newMVar : Expr) : MetaM Bool := do
let mvarType ← inferType mvar
forallBoundedTelescope mvarType numArgs fun xs _ => do
if xs.size != numArgs then return false
let some v ← mkLambdaFVarsWithLetDeps xs newMVar | return false
let some v ← checkAssignmentAux mvar.mvarId! #[] false v | return false
checkTypesAndAssign mvar v
-- See checkAssignment
partial def checkAssignmentAux (mvarId : MVarId) (fvars : Array Expr) (hasCtxLocals : Bool) (v : Expr) : MetaM (Option Expr) := do
run (check v) mvarId fvars hasCtxLocals v
partial def checkApp (e : Expr) : CheckAssignmentM Expr :=
e.withApp fun f args => do
let ctxMeta ← readThe Meta.Context
if f.isMVar && ctxMeta.config.ctxApprox && args.all Expr.isFVar then
let f ← check f
catchInternalId outOfScopeExceptionId
(do
let args ← args.mapM check
return mkAppN f args)
(fun ex => do
if !f.isMVar then
throw ex
if (← f.mvarId!.isDelayedAssigned) then
throw ex
let eType ← inferType e
let mvarType ← check eType
/- Create an auxiliary metavariable with a smaller context and "checked" type, assign `?f := fun _ => ?newMVar`
Note that `mvarType` may be different from `eType`. -/
let ctx ← read
let newMVar ← mkAuxMVar ctx.mvarDecl.lctx ctx.mvarDecl.localInstances mvarType
if (← assignToConstFun f args.size newMVar) then
pure newMVar
else
throw ex)
else
let f ← check f
let args ← args.mapM check
return mkAppN f args
partial def check (e : Expr) : CheckAssignmentM Expr := do
if !e.hasExprMVar && !e.hasFVar then
return e
else checkCache e fun _ =>
match e with
| .mdata _ b => return e.updateMData! (← check b)
| .proj _ _ s => return e.updateProj! (← check s)
| .lam _ d b _ => return e.updateLambdaE! (← check d) (← check b)
| .forallE _ d b _ => return e.updateForallE! (← check d) (← check b)
| .letE _ t v b _ => return e.updateLet! (← check t) (← check v) (← check b)
| .bvar .. => return e
| .sort .. => return e
| .const .. => return e
| .lit .. => return e
| .fvar .. => checkFVar e
| .mvar .. => checkMVar e
| .app .. =>
try
checkApp e
catch ex => match ex with
| .internal id =>
/-
If `ex` is an `CheckAssignmentM` internal exception and `e` is a beta-redex, we reduce `e` and try again.
This is useful for assignments such as `?m := (fun _ => A) a` where `a` is free variable that is not in
the scope of `?m`.
Note that, we do not try expensive reductions (e.g., `delta`). Thus, the following assignment
```lean
?m := Function.const 0 a
```
still fails because we do reduce the rhs to `0`. We assume this is not an issue in practice.
-/
if (id == outOfScopeExceptionId || id == checkAssignmentExceptionId) && e.isHeadBetaTarget then
checkApp e.headBeta
else
throw ex
| _ => throw ex
-- TODO: investigate whether the following feature is too expensive or not
/-
catchInternalIds [checkAssignmentExceptionId, outOfScopeExceptionId]
(checkApp e)
fun ex => do
let e' ← whnfR e
if e != e' then
check e'
else
throw ex
-/
end
end CheckAssignment
namespace CheckAssignmentQuick
partial def check
(hasCtxLocals : Bool)
(mctx : MetavarContext) (lctx : LocalContext) (mvarDecl : MetavarDecl) (mvarId : MVarId) (fvars : Array Expr) (e : Expr) : Bool :=
let rec visit (e : Expr) : Bool := Id.run do
if !e.hasExprMVar && !e.hasFVar then
return true
match e with
| .mdata _ b => visit b
| .proj _ _ s => visit s
| .app f a => visit f && visit a
| .lam _ d b _ => visit d && visit b
| .forallE _ d b _ => visit d && visit b
| .letE _ t v b _ => visit t && visit v && visit b
| .bvar .. | .sort .. | .const ..
| .lit .. => return true
| .fvar fvarId .. =>
if mvarDecl.lctx.contains fvarId then return true
if let some (LocalDecl.ldecl ..) := lctx.find? fvarId then
return false -- need expensive CheckAssignment.check
if fvars.any fun x => x.fvarId! == fvarId then
return true
return false -- We could throw an exception here, but we would have to use ExceptM. So, we let CheckAssignment.check do it
| .mvar mvarId' =>
let none := mctx.getExprAssignmentCore? mvarId' | return false -- use CheckAssignment.check to instantiate
if mvarId' == mvarId then return false -- occurs check failed, use CheckAssignment.check to throw exception
let some mvarDecl' := mctx.findDecl? mvarId' | return false
if hasCtxLocals then return false -- use CheckAssignment.check
if !mvarDecl'.lctx.isSubPrefixOf mvarDecl.lctx fvars then return false -- use CheckAssignment.check
let none := mctx.getDelayedMVarAssignmentCore? mvarId' | return false -- use CheckAssignment.check
return true
visit e
end CheckAssignmentQuick
/--
Auxiliary function used at `typeOccursCheckImp`.
Given `type`, it tries to eliminate "dependencies". For example, suppose we are trying to
perform the assignment `?m := f (?n a b)` where
```
?n : let k := g ?m; A -> h k ?m -> C
```
If we just perform occurs check `?m` at the type of `?n`, we get a failure, but
we claim these occurrences are ok because the type `?n a b : C`.
In the example above, `typeOccursCheckImp` invokes this function with `n := 2`.
Note that we avoid using `whnf` and `inferType` at `typeOccursCheckImp` to minimize the
performance impact of this extra check.
See test `typeOccursCheckIssue.lean` for an example where this refinement is needed.
The test is derived from a Mathlib file.
-/
private partial def skipAtMostNumBinders (type : Expr) (n : Nat) : Expr :=
match type, n with
| .forallE _ _ b _, n+1 => skipAtMostNumBinders b n
| .mdata _ b, n => skipAtMostNumBinders b n
| .letE _ _ v b _, n => skipAtMostNumBinders (b.instantiate1 v) n
| type, _ => type
/-- `typeOccursCheck` implementation using unsafe (i.e., pointer equality) features. -/
private unsafe def typeOccursCheckImp (mctx : MetavarContext) (mvarId : MVarId) (v : Expr) : Bool :=
if v.hasExprMVar then
visit v |>.run' mkPtrSet
else
true
where
alreadyVisited (e : Expr) : StateM (PtrSet Expr) Bool := do
if (← get).contains e then
return true
else
modify fun s => s.insert e
return false
occursCheck (type : Expr) : Bool :=
let go : StateM MetavarContext Bool := do
Lean.occursCheck mvarId type
-- Remark: it is ok to discard the the "updated" `MetavarContext` because
-- this function assumes all assigned metavariables have already been
-- instantiated.
go.run' mctx
visitMVar (mvarId' : MVarId) (numArgs : Nat := 0) : Bool :=
if let some mvarDecl := mctx.findDecl? mvarId' then
occursCheck (skipAtMostNumBinders mvarDecl.type numArgs)
else
false
visitApp (e : Expr) : StateM (PtrSet Expr) Bool :=
e.withApp fun f args => do
unless (← args.allM visit) do
return false
if f.isMVar then
return visitMVar f.mvarId! args.size
else
visit f
visit (e : Expr) : StateM (PtrSet Expr) Bool := do
if !e.hasExprMVar then
return true
else if (← alreadyVisited e) then
return true
else match e with
| .mdata _ b => visit b
| .proj _ _ s => visit s
| .app .. => visitApp e
| .lam _ d b _ => visit d <&&> visit b
| .forallE _ d b _ => visit d <&&> visit b
| .letE _ t v b _ => visit t <&&> visit v <&&> visit b
| .mvar mvarId' => return visitMVar mvarId'