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/Logic/L05_Apply.lean

45 lines
1.3 KiB
Plaintext

2 years ago
import TestGame.Metadata
2 years ago
import Mathlib
2 years ago
Game "TestGame"
World "Logic"
2 years ago
Level 6
Title "Implikation"
Introduction
"Als nächstes widmen wir uns der Implikation $A \\Rightarrow B$.
2 years ago
Mit zwei logischen Aussagen `(A : Prop)`, `(B : Prop)` schreibt man eine Implikation
als `A → B` (`\\to`).
2 years ago
Wenn man als Goal $B$ hat und eine Impikation `(g : A → B)` kann man mit `apply g` diese
anwenden, worauf das Goal $A$ ist. Auf Papier würde man an der Stelle
folgendes zu schreiben: \"Es genügt $A$ zu zeigen, denn $A$ impliziert $B$\".
*Bemerke:* Das ist der selbe Pfeil, der später auch für Funktionen wie `` gebraucht
wird, deshalb heisst er `\\to`."
2 years ago
Statement
"Seien $A$, $B$ logische Aussagen, wobei $A$ wahr ist und $A$ impliziert $B$.
Zeige, dass $B$ wahr ist."
2 years ago
(A B : Prop) (hA : A) (g : A → B) : B := by
apply g
assumption
Hint (A : Prop) (B : Prop) (hA : A) (g : A → B) : A =>
"Mit `apply g` kannst du die Implikation `g` anwenden."
Hint (A : Prop) (B : Prop) (hA : A) (g : A → B) : A =>
"Nachdem du die Implikation `A → B` angewendet hast, musst du nur noch $A$ zeigen,
2 years ago
dafür hast du bereits einen Beweis in den Annahmen."
2 years ago
Tactics apply assumption
-- Mathjax notes
-- `\\\\( \\\\)` or `$ $` for inline
-- and `$$ $$` block.
-- align block:
-- $$\\begin{aligned} 2x - 4 &= 6 \\\\\\\\ 2x &= 10 \\\\\\\\ x &= 5 \\end{aligned}$$