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/Negation/L02_Contra.lean

30 lines
676 B
Plaintext

2 years ago
import TestGame.Metadata
import Std.Tactic.RCases
import Mathlib.Tactic.LeftRight
Game "TestGame"
World "Contradiction"
Level 2
Title "Widerspruch"
Introduction
"
Ähnlich siehts aus, wenn man Annahmen hat, die direkte Negierung voneinander sind,
also `(h : A)` und `(g : ¬ A)`. (`\\not`)
"
Statement
"Ein Widerspruch impliziert alles."
(A : Prop) (n : ) (h : ∃ x, 2 * x = n) (g : ¬ (∃ x, 2 * x = n)) : A := by
contradiction
Conclusion
"
Detail: `¬ A` ist übrigens als `A → false` implementiert, was aussagt, dass
\"falls `A` wahr ist, impliziert das `false` und damit einen Widerspruch\".
"
-- TODO: Oder doch ganz entfernen?
Tactics contradiction