You signed in with another tab or window. Reload to refresh your session.You signed out in another tab or window. Reload to refresh your session.You switched accounts on another tab or window. Reload to refresh your session.Dismiss alert
The definitional equality check at the meta level (not the kernel) always fails when checking equality of two expressions that are both references to the same constant when the universe parameters are not equal. However, it is still possible that both might unfold to definitionally equal terms or that they may satisfy the rule that all inhabitants of unit-like types are definitionally equal.
Steps to Reproduce
Use the following code:
abbrev T.{u} : Unit := (fun (α : Sort u) => ()) PUnit.{u}
set_option pp.universes true
defunitUnique (x y : Unit) : x = y := by rfl
deftestUnique.{u, v} : T.{u} = T.{v} := unitUnique _ _
/--error: type mismatch rfl.{1}has type Eq.{1} T.{u} T.{u} : Propbut is expected to have type Eq.{1} T.{u} T.{v} : Prop-/
#guard_msgs indeftest.{u, v} : T.{u} = T.{v} := rfl
deftest'.{u, v} : T.{u} = T.{v} := rfl (a := ())
theoremtest''.{u, v} : T.{u} = T.{v} := by
unfold T
rfl
-- Use kernel defeqdeftest'''.{u, v} : T.{u} = T.{v} := by
run_tac do
Lean.Elab.Tactic.closeMainGoalUsing `test''' fun t _ =>
Lean.Meta.mkEqRefl t.eq?.get!.2.1/--info: def test'''.{u, v} : Eq.{1} T.{u} T.{v} :=Eq.refl.{1} T.{u}-/
#guard_msgs in#print test'''
Expected behavior:
test should type check successfully, just like the other two versions do. There should be a delta reduction of T, as the kernel is using, or alternatively an application of the rule for unit-like types.
Actual behavior:
It does not type check.
There seems to be a general assumption in this code that all constants are injective WRT universe levels. I think this is true for type constructors and constructors, but not for definitions. This assumption is manifest in the following two locations:
This PR fixes a bug at the definitional equality test (`isDefEq`). At
unification constraints of the form `c.{u} =?= c.{v}`, it was not
trying to unfold `c`. This bug did not affect the kernel.
closes#6117
This PR fixes a bug at the definitional equality test (`isDefEq`). At
unification constraints of the form `c.{u} =?= c.{v}`, it was not trying
to unfold `c`. This bug did not affect the kernel.
closes#6117
Description
The definitional equality check at the meta level (not the kernel) always fails when checking equality of two expressions that are both references to the same constant when the universe parameters are not equal. However, it is still possible that both might unfold to definitionally equal terms or that they may satisfy the rule that all inhabitants of unit-like types are definitionally equal.
Steps to Reproduce
Use the following code:
Expected behavior:
test
should type check successfully, just like the other two versions do. There should be a delta reduction ofT
, as the kernel is using, or alternatively an application of the rule for unit-like types.Actual behavior:
It does not type check.
There seems to be a general assumption in this code that all constants are injective WRT universe levels. I think this is true for type constructors and constructors, but not for definitions. This assumption is manifest in the following two locations:
lean4/src/Lean/Meta/ExprDefEq.lean
Line 2058 in 5a99cb3
lean4/src/Lean/Meta/ExprDefEq.lean
Line 1390 in 5a99cb3
Versions
This occurs at least in 4.12 and the latest nightly (
4.15.0-nightly-2024-11-15
)Link to live.lean-lang.org
Additional Information
For applications, the problem is not manifest. All of the following are accepted:
Impact
Add 👍 to issues you consider important. If others are impacted by this issue, please ask them to add 👍 to it.
The text was updated successfully, but these errors were encountered: