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/LeftOvers/L09_Or.lean

71 lines
2.2 KiB
Plaintext

import Adam.Metadata
2 years ago
import Std.Tactic.RCases
import Mathlib.Tactic.LeftRight
set_option tactic.hygienic false
Game "Adam"
World "Implication"
Level 9
2 years ago
Title "Oder"
Introduction
"
Übung macht den Meister... Benutze alle vier Methoden mit UND und ODER
umzugehen um folgende Aussage zu beweisen.
| | Und | Oder |
|---------|:-------------------------|:--------------------------|
| Annahme | `rcases h with ⟨h₁, h₂⟩` | `rcases h with h₁ \\| h₂` |
| Goal | `constructor` | `left`/`right` |
2 years ago
"
Statement and_or_imp
"Angenommen $(A \\land B) \\lor (A \\Rightarrow C)$ und $A$ sind wahr, zeige dass
$B \\lor (C \\land A)$ wahr ist."
2 years ago
(A B C : Prop) (h : (A ∧ B) (A → C)) (hA : A) : (B (C ∧ A)) := by
rcases h with h₁ | h₂
left
rcases h₁ with ⟨x, y⟩
assumption
right
constructor
apply h₂
assumption
assumption
HiddenHint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B (A → C)) (hA : A) : B (C ∧ A) =>
"Ein ODER in den Annahmen teilt man mit `rcases h with h₁ | h₂`."
-- If starting with `left`.
Hint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B (A → C)) : B =>
"Da kommst du nicht mehr weiter..."
-- If starting with `right`.
Hint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B (A → C)) : (C ∧ A) =>
"Da kommst du nicht mehr weiter..."
HiddenHint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B) (hA : A) : B (C ∧ A) =>
"`left` oder `right`?"
HiddenHint (A : Prop) (B : Prop) (C : Prop) (h : B) (hA : A) : B (C ∧ A) =>
"`left` oder `right`?"
2 years ago
Hint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B) (hA : A) : B (C ∧ A) =>
"Ein UND in den Annahmen kann man mit `rcases h with ⟨h₁, h₂⟩` aufteilen."
2 years ago
Hint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B) (hA : A) : B =>
"Ein UND in den Annahmen kann man mit `rcases h with ⟨h₁, h₂⟩` aufteilen."
2 years ago
Hint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B) : C =>
2 years ago
"Sackgasse."
Hint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B) : C ∧ A =>
2 years ago
"Hmmm..."
Hint (A : Prop) (B : Prop) (C : Prop) (h : A → C) : C ∧ A =>
2 years ago
"Ein UND im Goal kann mit `constructor` aufgeteilt werden."
NewTactic left right assumption constructor rcases