Files
NNG/Game/Levels/Implication/L09zero_ne_succ.lean
2023-10-11 17:53:51 +01:00

42 lines
1.2 KiB
Lean4
Raw Blame History

This file contains ambiguous Unicode characters
This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.
import Game.Levels.Implication.L08ne
World "Implication"
Level 9
Title "zero_ne_succ"
LemmaTab "Peano"
namespace MyNat
LemmaDoc MyNat.zero_ne_succ as "zero_ne_succ" in "Peano" "
`zero_ne_succ n` is the proof that `0 ≠ succ n`.
In Lean, `a ≠ b` is *defined to mean* `a = b → False`. Hence
`zero_ne_succ n` is really a proof of `0 = succ n → False`.
Here `False` is a generic false statement. This means that
you can `apply zero_ne_succ at h` if `h` is a proof of `0 = succ n`.
"
Introduction "
As warm-up for `2 + 2 ≠ 5` let's prove `0 ≠ 1`. To do this we need to
introduce Peano's last axiom `zero_ne_succ n`, a proof that `0 ≠ succ n`.
To learn about this result, click on it in the list of lemmas on the right.
"
LemmaDoc MyNat.zero_ne_one as "zero_ne_one" in "numerals" "
`zero_ne_one` is a proof of `0 ≠ 1`.
"
NewLemma MyNat.zero_ne_succ MyNat.zero_ne_one
/-- $0\neq1$. -/
Statement zero_ne_one : (0 : ) 1 := by
Hint "Start with `intro h`."
intro h
rw [one_eq_succ_zero] at h -- **TODO** this line is not needed :-/
apply zero_ne_succ at h -- **TODO** cripple `apply`.
exact h
Conclusion "Remember, `x ≠ y` is *notation* for `x = y → false`"