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/adam/Adam/Levels/Implication/L12_Rw.lean

44 lines
1.2 KiB
Plaintext

import Adam.Metadata
2 years ago
import Std.Tactic.RCases
import Mathlib.Tactic.Cases
import Mathlib.Logic.Basic
Game "Adam"
2 years ago
World "Implication"
Level 12
Title "Lemmas"
Introduction
"
**Operationsleiter**: Wieder etwas für den Kollegen …. Und er wollte wieder einen Beweise ohne `apply`. Ich sehe hier auch, dass ich mir schon einmal etwas hierzu notiert hatte. Richtig, es gibt da dieses Lemma:
2 years ago
```
lemma not_not (A : Prop) : ¬¬A ↔ A
```
**Operationsleiter**: Schafft Ihr das damit?
2 years ago
"
Statement (A B C : Prop) : (A ∧ (¬¬C)) (¬¬B) ∧ C ↔ (A ∧ C) B ∧ (¬¬C) := by
Hint "**Robo**: Ein Lemma, das wie `not_not` ein `↔` oder `=` im Statement hat, kann
auch mit `rw [not_not]` verwendet werden."
rw [not_not]
Hint "**Du**: Häh, wieso hat das jetzt 2 von 3 der `¬¬` umgeschrieben?
**Robo**: `rw` schreibt nur das erste um, das es findet, also `¬¬C`. Aber weil dieses
mehrmals vorkommt, werden die alle ersetzt …
2 years ago
**Du**: Ah, und `¬¬B` ist etwas anderes, also brauche ich das Lemma nochmals."
2 years ago
rw [not_not]
Conclusion
"
**Du**: Wir sind schon fertig …?
2 years ago
**Robo**: Ja, `rw` versucht immer anschließend `rfl` aufzurufen, und das hat hier funktioniert.
2 years ago
"
DisabledTactic tauto apply
NewLemma not_not
LemmaTab "Logic"