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/L04_Contra.lean

32 lines
925 B
Plaintext

2 years ago
import TestGame.Metadata
import Std.Tactic.RCases
import Mathlib.Tactic.LeftRight
2 years ago
Game "TestGame"
World "Contradiction"
Level 4
Title "Ad absurdum"
2 years ago
Introduction
"
Im weiteren kann man auch Widersprüche erhalten, wenn man Annahmen der Form
`A = B` hat, wo Lean weiss, dass `A und `B` unterschiedlich sind, z.B. `0 = 1` in ``
oder auch Annahmen der Form `A ≠ A` (`\\ne`).
2 years ago
"
Statement
"Ein Widerspruch impliziert alles."
(A : Prop) (a b c : ) (g₁ : a = b) (g₂ : b = c) (h : a ≠ c) : A := by
rw [g₁] at h
contradiction
Message (A : Prop) (a : ) (b : ) (c : ) (g₁ : a = b) (g₂ : b = c) (h : a ≠ c) : A =>
"Recap: `rw [...] at h` hilft dir hier."
2 years ago
Message (A : Prop) (a : ) (b : ) (c : ) (g₁ : a = b) (g₂ : b = c) (h : b ≠ c) : A =>
"`b ≠ c` muss man als `¬ (b = c)` lesen. Deshalb findet `contradiction` hier direkt
einen Widerspruch."
2 years ago
Tactics contradiction