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.
|
|
|
|
import Adam.Metadata
|
|
|
|
|
import Mathlib.Tactic.Ring
|
|
|
|
|
|
|
|
|
|
--set_option tactic.hygienic false
|
|
|
|
|
|
|
|
|
|
Game "Adam"
|
|
|
|
|
World "Predicate"
|
|
|
|
|
Level 1
|
|
|
|
|
|
|
|
|
|
Title "Natürliche Zahlen"
|
|
|
|
|
|
|
|
|
|
Introduction
|
|
|
|
|
""
|
|
|
|
|
|
|
|
|
|
Statement (x y : ℕ) : (x + y) ^ 2 = x ^ 2 + 2 * x * y + y ^ 2 := by
|
|
|
|
|
Hint "**Du**: Das ist doch Schulmathematik! Man rechnet das einfach aus,
|
|
|
|
|
indem man die Terme umsortiert.
|
|
|
|
|
|
|
|
|
|
**Robo**: Wenn die Gleichung stimmt, kannst Du auf Leansch sogar einfach mit `ring` beweisen, dass das so ist.
|
|
|
|
|
|
|
|
|
|
**Du**: Aber `ℕ` ist doch gar kein Ring?
|
|
|
|
|
|
|
|
|
|
**Robo**: `ring` funktioniert sogar für sogenannte Halbringe. Ich glaube, man sagt `ring`, weil es in (kommutativen) Ringen am besten funktioniert.
|
|
|
|
|
"
|
|
|
|
|
ring
|
|
|
|
|
|
|
|
|
|
Conclusion
|
|
|
|
|
""
|
|
|
|
|
|
|
|
|
|
NewTactic ring
|