You cannot select more than 25 topics Topics must start with a letter or number, can include dashes ('-') and can be up to 35 characters long.
lean4game/server/testgame/TestGame/Levels/Level2.lean

39 lines
1.2 KiB
Plaintext

This file contains ambiguous Unicode characters!

This file contains ambiguous Unicode characters that may be confused with others in your current locale. If your use case is intentional and legitimate, you can safely ignore this warning. Use the Escape button to highlight these characters.

import TestGame.Metadata
Game "TestGame"
World "Old"
Level 2
Title "The rewriting spell"
Introduction
"
The rewrite spell is the way to \"substitute in\" the value
of an expression. In general, if you have a hypothesis of the form `A = B`, and your
goal mentions the left hand side `A` somewhere, then
the `rewrite` tactic will replace the `A` in your goal with a `B`.
The documentation for `rewrite` just appeared in your spell book.
Play around with the menus and see what is there currently.
More information will appear as you progress.
Take a look in the top right box at what we have.
The variables $x$ and $y$ are natural numbers, and we have
an assumption `h` that $y = x + 7$. Our goal
is to prove that $2y=2(x+7)$. This goal is obvious -- we just
substitute in $y = x+7$ and we're done. In Lean, we do
this substitution using the `rewrite` spell. This spell takes a list of equalities
or equivalences so you can cast `rewrite [h]`.
"
Statement "" (x y : ) (h : y = x + 7) : 2 * y = 2 * (x + 7) := by
rewrite [h]
rfl
Message (x : ) (y : ) (h : y = x + 7) : 2*(x + 7) = 2*(x + 7) =>
"Great! Now the goal should be easy to reach using the `rfl` spell."
Conclusion "Congratulations for completing your second level!"
Tactics rfl rewrite