-
Notifications
You must be signed in to change notification settings - Fork 465
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
fix: use maxType when building expression in expression tree elaborat…
…or (#4215) The expression tree elaborator computes a "maxType" that every leaf term can be coerced to, but the elaborator was not ensuring that the entire expression tree would have maxType as its type. This led to unexpected errors in examples such as ```lean example (a : Nat) (b : Int) : a = id (a * b^2) := sorry ``` where it would say it could not synthesize an `HMul Int Int Nat` instance (the `Nat` would propagate from the `a` on the LHS of the equality). The issue in this case is that `HPow` uses default instances, so while the expression tree elaborator decides that `a * b^2` should be referring to an `Int`, the actual elaborated type is temporarily a metavariable. Then, when the binrel elaborator is looking at both sides of the equality, it decides that `Nat` will work and coercions don't need to be inserted. The fix is to unify the type of the resulting elaborated expression with the computed maxType. One wrinkle is that `hasUncomparable` being false is a valid test only if there are no leaf terms with unknown types (if they become known, it could change `hasUncomparable` to true), so this unification is only performed if the leaf terms all have known types. Fixes issue described by Floris van Doorn on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/elaboration.20issue.20involving.20powers.20and.20sums/near/439243587).
- Loading branch information
Showing
4 changed files
with
58 additions
and
22 deletions.
There are no files selected for viewing
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file was deleted.
Oops, something went wrong.
This file was deleted.
Oops, something went wrong.
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,41 @@ | ||
/-! | ||
# Tests for the expression tree elaborator (`binop%`, etc.) | ||
-/ | ||
|
||
/-! | ||
Some basic Int/Nat examples | ||
-/ | ||
|
||
example (n : Nat) (i : Int) : n + i = i + n := by | ||
rw [Int.add_comm] | ||
|
||
def f1 (a : Int) (b c : Nat) : Int := | ||
a + (b - c) | ||
|
||
def f2 (a : Int) (b c : Nat) : Int := | ||
(b - c) + a | ||
|
||
/-- | ||
info: def f1 : Int → Nat → Nat → Int := | ||
fun a b c => a + (↑b - ↑c) | ||
-/ | ||
#guard_msgs in | ||
#print f1 | ||
|
||
/-- | ||
info: def f2 : Int → Nat → Nat → Int := | ||
fun a b c => ↑b - ↑c + a | ||
-/ | ||
#guard_msgs in | ||
#print f2 | ||
|
||
|
||
/-! | ||
Interaction with default instances for pow. This used to fail with not being able | ||
to synthesize an `HMul Int Int Nat` instance because the type of | ||
the result of `*` wasn't being set to `Int`. | ||
-/ | ||
|
||
/-- info: ∀ (a : Nat) (b : Int), ↑a = id (↑a * b ^ 2) : Prop -/ | ||
#guard_msgs in | ||
#check ∀ (a : Nat) (b : Int), a = id (a * b^2) |