Files
NNG/Game/Levels/Tutorial/L06add_zero2.lean
Kevin Buzzard d80a520359 tinkering
2023-10-22 15:50:20 +01:00

37 lines
873 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.Metadata
import Game.MyNat.Addition
World "Tutorial"
Level 6
Title "Precision rewriting"
Introduction
"
## Precision rewriting
In the last level, there was `b + 0` and `c + 0`,
and `rw [add_zero]` changed the first one it saw,
which was `b + 0`. Let's learn how to tell Lean
to change `c + 0` first by giving `add_zero` an
explicit input.
"
namespace MyNat
/-- $a+(b+0)+(c+0)=a+b+c.$ -/
Statement (a b c : ) : a + (b + 0) + (c + 0) = a + b + c := by
Hint "Try `rw [add_zero c]`."
rw [add_zero c]
Hint "`add_zero c` is a proof of `c + 0 = c` so that was what got rewritten.
You can now change `b + 0` to `b` with `rw [add_zero]` or `rw [add_zero b]`. You
can usually stick to `add_zero` unless you need real precision."
rw [add_zero]
rfl
LemmaTab "Add"
Conclusion "
Let's now learn about Peano's second axiom for addition, `add_succ`.
"