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
induction will use instances on the original variable in function applications on the new variable. The tactic will succeed, and indeed create new goals, but the proof will have a type error.
/--warning: declaration uses 'sorry'---error: application type mismatch @Fin.instOfNat n inst✝argument has type NeZero n✝but function has type [inst : NeZero n] → {i : Nat} → OfNat (Fin n) i-/
#guard_msgs inexample (n: Nat) [NeZero n] : (0 : Fin n) = 0 := by
induction n <;> sorry
Expected behavior: An error on the tactic as it should try to synthesize NeZero 0 and NeZero n (where the original n is succ of this n) and both should fail.
Actual behavior: The tactic uses the original NeZero n✝ instance to synthesize Fin.instOfNat, for the new (0 : Fin n), which is not type-correct. Yet the tactic does not fail.
Prerequisites
Please put an X between the brackets as you perform the following steps:
https://github.com/leanprover/lean4/issues
Avoid dependencies to Mathlib or Batteries.
https://live.lean-lang.org/#project=lean-nightly
(You can also use the settings there to switch to “Lean nightly”)
Description
induction
will use instances on the original variable in function applications on the new variable. The tactic will succeed, and indeed create new goals, but the proof will have a type error.Context
This was discovered during usage of https://github.com/leanprover-community/repl, where we found a successful sequence of tactics resulting in an unsuccessful proof.
Steps to Reproduce
Expected behavior: An error on the tactic as it should try to synthesize
NeZero 0
andNeZero n
(where the originaln
issucc
of thisn
) and both should fail.Actual behavior: The tactic uses the original
NeZero n✝
instance to synthesizeFin.instOfNat
, for the new(0 : Fin n)
, which is not type-correct. Yet the tactic does not fail.Versions
Lean 4.16.0-rc1
Target: x86_64-unknown-linux-gnu
live.lean-lang.org
Additional Information
Initially, #4246 looked similar, but I think it's not the same.
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: