Files
NNG/Game/Levels/AdvAddition/Level_13.lean
2023-08-11 15:06:01 +02:00

39 lines
866 B
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.AdvAddition.Level_12
World "AdvAddition"
Level 13
Title "ne_succ_self"
open MyNat
Introduction
"
The last level in Advanced Addition World is the statement
that $n\\not=\\operatorname{succ}(n)$.
"
/-- For any natural number $n$, we have
$$ n \\neq \\operatorname{succ}(n). $$ -/
Statement --ne_succ_self
(n : ) : n succ n := by
Hint (hidden := true) "I would start a proof by induction on `n`."
induction n with d hd
· apply zero_ne_succ
· Hint (hidden := true) "If you have no clue, you could start with `rw [Ne, Not]`."
Branch
rw [Ne, Not]
intro hs
apply hd
apply succ_inj
exact hs
LemmaTab "Nat"
Conclusion
"
Congratulations. You've completed Advanced Addition World and can move on
to Advanced Multiplication World (after first doing
Multiplication World, if you didn't do it already).
"