Files
NNG/Game/Levels/Implication/L04succ_inj.lean

72 lines
2.4 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.L03apply
World "Implication"
Level 4
Title "succ_inj : the successor function is injective"
namespace MyNat
LemmaTab "Peano"
Introduction
"
If `a` and `b` are numbers, then `succ_inj a b` is a proof
that `succ a = succ b` implies `a = b`. Click on this theorem in the *Peano*
tab for more information.
Peano had this theorem as an axiom, but in Functional Programming World
we will show how to prove it in Lean. Right now let's just assume it,
and let's prove $x+1=4 \\implies x=3$ using it. Again, we will proceed
by manipulating our hypothesis until it becomes the goal. I will
walk you through this level.
"
LemmaDoc MyNat.succ_inj as "succ_inj" in "Peano" "
# Statement
If $a$ and $b$ are numbers, then
`succ_inj a b` is the proof that
$ (\\operatorname{succ}(a) = \\operatorname{succ}(b))\\implies a=b$.
## More technical details
There are other ways to think about `succ_inj`.
You can think about `succ_inj` itself as a function which takes two
numbers $$a$$ and $$b$$ as input, and outputs a proof of
$ (\\operatorname{succ}(a) = \\operatorname{succ}(b))\\implies a=b$.
You can think of `succ_inj` itself as a proof; it is the proof
that `succ` is an injective function. In other words,
`succ_inj` is a proof of
$\\forall a, b\\in \\mathbb{N}, (\\operatorname{succ}(a) = \\operatorname{succ}(b))\\implies a=b$.
`succ_inj` was postulated as an axiom by Peano, but
in Lean it can be proved using `pred`, a mathematically
pathological function.
"
NewLemma MyNat.succ_inj
/-- If $x+1=4$ then $x=3$. -/
Statement (x : ) (h : x + 1 = 4) : x = 3 := by
Hint "Let's first get `h` into the form `succ x = succ 3` so we can
apply `succ_inj`. First rewrite
`four_eq_succ_three` at `h` to change the 4 on the right hand side."
rw [four_eq_succ_three] at h
Hint "Now rewrite `succ_eq_add_one` backwards at `h`
to get the right hand side."
Hint (hidden := true) "`rw [← succ_eq_add_one] at h`. Look at the
docs for `rw` for an explanation. Type `←` with `\\l`."
rw [succ_eq_add_one] at h
Hint "Now let's `apply` our new theorem. Execute `apply succ_inj at h`
to change `h` to a proof of `x = 3`."
apply succ_inj at h
Hint "Now finish in one line."
Hint (hidden := true) "And now we've deduced what we wanted to prove: the goal is one of our assumptions.
Finish the level with `exact h`."
exact h
Conclusion "In the next level, we'll do the same proof but backwards."