rename [New/Only/Disabled][Tactic/Lemma/Definition]

pull/54/head
Jon Eugster 2 years ago
parent 01199c3793
commit 93d05c9703

@ -201,19 +201,19 @@ elab "TacticDoc" name:ident content:str : command =>
content := content.getString })
/-- Declare tactics that are introduced by this level. -/
elab "NewTactics" args:ident* : command => do
elab "NewTactic" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Tactic name
modifyCurLevel fun level => pure {level with tactics := {level.tactics with new := names}}
/-- Declare tactics that are temporarily disabled in this level -/
elab "DisabledTactics" args:ident* : command => do
elab "DisabledTactic" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Tactic name
modifyCurLevel fun level => pure {level with tactics := {level.tactics with disabled := names}}
/-- Temporarily disable all tactics except the ones declared here -/
elab "OnlyTactics" args:ident* : command => do
elab "OnlyTactic" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Tactic name
modifyCurLevel fun level => pure {level with tactics := {level.tactics with only := names}}
@ -231,19 +231,19 @@ elab "DefinitionDoc" name:ident content:str : command =>
content := content.getString })
/-- Declare definitions that are introduced by this level. -/
elab "NewDefinitions" args:ident* : command => do
elab "NewDefinition" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Definition name
modifyCurLevel fun level => pure {level with definitions := {level.definitions with new := names}}
/-- Declare definitions that are temporarily disabled in this level -/
elab "DisabledDefinitions" args:ident* : command => do
elab "DisabledDefinition" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Definition name
modifyCurLevel fun level => pure {level with definitions := {level.definitions with disabled := names}}
/-- Temporarily disable all definitions except the ones declared here -/
elab "OnlyDefinitions" args:ident* : command => do
elab "OnlyDefinition" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Definition name
modifyCurLevel fun level => pure {level with definitions := {level.definitions with only := names}}
@ -264,19 +264,19 @@ elab "LemmaDoc" name:ident "as" userName:ident "in" category:str content:str : c
content := content.getString })
/-- Declare lemmas that are introduced by this level. -/
elab "NewLemmas" args:ident* : command => do
elab "NewLemma" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Lemma name
modifyCurLevel fun level => pure {level with lemmas := {level.lemmas with new := names}}
/-- Declare lemmas that are temporarily disabled in this level -/
elab "DisabledLemmas" args:ident* : command => do
elab "DisabledLemma" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Lemma name
modifyCurLevel fun level => pure {level with lemmas := {level.lemmas with disabled := names}}
/-- Temporarily disable all lemmas except the ones declared here -/
elab "OnlyLemmas" args:ident* : command => do
elab "OnlyLemma" args:ident* : command => do
let names := args.map (·.getId)
for name in names do checkInventoryDoc .Lemma name
modifyCurLevel fun level => pure {level with lemmas := {level.lemmas with only := names}}

@ -63,8 +63,8 @@ have k : ¬ B
Conclusion ""
NewTactics «have»
DisabledTactics «suffices»
NewTactic «have»
DisabledTactic «suffices»
NewDefinitions Even Odd
NewLemmas not_even not_odd
NewDefinition Even Odd
NewLemma not_even not_odd

@ -60,7 +60,7 @@ Hint (A : Prop) (B : Prop) (h : A → ¬ B) (k : A) (f : B) : False =>
Conclusion ""
NewTactics «suffices»
DisabledTactics «have»
NewDefinitions Even Odd
NewLemmas not_even not_odd
NewTactic «suffices»
DisabledTactic «have»
NewDefinition Even Odd
NewLemma not_even not_odd

@ -52,4 +52,4 @@ HiddenHint (A : Prop) (B : Prop) (h : A → B) (b : ¬ B) (a : A) : False =>
Conclusion ""
NewTactics by_contra contradiction apply assumption
NewTactic by_contra contradiction apply assumption

@ -46,4 +46,4 @@ HiddenHint (A : Prop) (B : Prop) : A → B ↔ (¬ B → ¬ A) =>
Conclusion ""
NewTactics contradiction constructor intro by_contra apply assumption
NewTactic contradiction constructor intro by_contra apply assumption

@ -61,6 +61,6 @@ Hint (n : ) : Even n → ¬Odd (n ^ 2) =>
Hint (n : ) : Even n → Even (n ^ 2) =>
"Diese Aussage hast du bereits als Lemma bewiesen, schau mal in der Bibliothek."
NewTactics contrapose rw apply
NewDefinitions Even Odd
NewLemmas not_even not_odd even_square
NewTactic contrapose rw apply
NewDefinition Even Odd
NewLemma not_even not_odd even_square

@ -71,6 +71,6 @@ HiddenHint (n: ) (h : Odd (n ^ 2)) (g : Even n) : Even (n ^ 2) =>
"Probiers mit `apply ...`"
NewTactics contradiction by_contra rw apply assumption -- TODO: suffices, have
NewDefinitions Even Odd
NewLemmas not_odd not_even even_square
NewTactic contradiction by_contra rw apply assumption -- TODO: suffices, have
NewDefinition Even Odd
NewLemma not_odd not_even even_square

@ -58,8 +58,8 @@ show that $f(x) = x^2 + 2x + 1$.
unfold f
ring
NewTactics «let»
OnlyTactics «let» intro unfold ring
NewTactic «let»
OnlyTactic «let» intro unfold ring
HiddenHint : ∀ x, f x = x ^ 2 + 2 * x + 1 =>
"Fang zuerst wie immer mit `intro x` an."

@ -19,4 +19,4 @@ Statement
(U T V : Type _) (f : U → V) (g : V → T) (x : U) : (g ∘ f) x = g (f x) := by
simp only [Function.comp_apply]
NewLemmas Function.comp_apply
NewLemma Function.comp_apply

@ -57,9 +57,9 @@ Statement ""
unfold f
ring
NewTactics funext by_cases simp_rw linarith
NewTactic funext by_cases simp_rw linarith
NewLemmas not_le if_pos if_neg
NewLemma not_le if_pos if_neg
Hint : f ∘ g = g ∘ f =>

@ -29,8 +29,8 @@ Statement "" : Injective (fun (n : ) ↦ n^3 + (n + 3)) := by
intro a b
simp
NewDefinitions Injective
NewLemmas StrictMono.injective StrictMono.add Odd.strictMono_pow
NewDefinition Injective
NewLemma StrictMono.injective StrictMono.add Odd.strictMono_pow
Hint : Injective fun (n : ) => n ^ 3 + (n + 3) =>

@ -22,7 +22,7 @@ Statement "" : Injective (fun (n : ) ↦ n + 3) := by
intro a b
simp
NewDefinitions Injective
NewDefinition Injective
Hint : Injective (fun (n : ) ↦ n + 3) =>
"**Robo**: `Injective` ist als `∀ \{a b : U}, f a = f b → a = b`

@ -37,4 +37,4 @@ HiddenHint (A : Prop) (B : Prop) (hb : B) : A → (A ∧ B) =>
Hint (A : Prop) (B : Prop) (ha : A) (hb : B) : (A ∧ B) =>
"Jetzt kannst du die Taktiken aus dem letzten Kapitel verwenden."
NewTactics intro constructor assumption
NewTactic intro constructor assumption

@ -39,4 +39,4 @@ dass $B$ wahr ist."
HiddenHint (A : Prop) (B : Prop) (ha : A) (h : A → B): B =>
"Mit `revert ha` kann man die Annahme `ha` als Implikationsprämisse vorne ans Goal anhängen."
NewTactics revert assumption
NewTactic revert assumption

@ -32,7 +32,7 @@ HiddenHint (A : Prop) (B : Prop) (hA : A) (g : A → B) : A =>
"Nachdem du die Implikation `A → B` angewendet hast, musst du nur noch $A$ zeigen,
dafür hast du bereits einen Beweis in den Annahmen."
NewTactics apply assumption
NewTactic apply assumption
-- Katex notes
-- `\\( \\)` or `$ $` for inline

@ -32,4 +32,4 @@ HiddenHint (A : Prop) (B : Prop) (C : Prop) (hA : A) (f : A → B) (g : B → C)
"Du willst $C$ beweisen. Suche also nach einer Implikation $\\ldots \\Rightarrow C$ und wende
diese mit `apply` an."
NewTactics intro apply assumption revert
NewTactic intro apply assumption revert

@ -48,6 +48,6 @@ Hint (A : Prop) (B : Prop) (C : Prop) (D : Prop) (E : Prop) (F : Prop)
(i : D → E) (k : E → F) (m : C → F) : D =>
"Sackgasse. Probier doch einen anderen Weg."
NewTactics apply assumption revert intro
NewTactic apply assumption revert intro
-- https://www.jmilne.org/not/Mamscd.pdf

@ -38,6 +38,6 @@ aus mehreren Einzelteilen besteht: `⟨A → B, B → A⟩`. Man sagt also Lean,
ob das Goal aus solchen Einzelteilen \"konstruiert\" werden kann.
"
NewTactics constructor assumption
NewTactic constructor assumption
-- TODO : `case mpr =>` ist mathematisch noch sinnvoll.

@ -71,4 +71,4 @@ Hint (A : Prop) (B : Prop) (C : Prop) (D : Prop) (h₁ : C ↔ D) (h₂ : D ↔
"Das ist nicht der optimale Weg..."
NewTactics rw assumption
NewTactic rw assumption

@ -41,4 +41,4 @@ Hint (A : Prop) (B : Prop) (C : Prop) (h : A ↔ B) (g : B → C) (hA : A) : B =
Conclusion "Im nächsten Level findest du die zweite Option."
NewTactics apply assumption
NewTactic apply assumption

@ -36,7 +36,7 @@ Conclusion
"
"
NewTactics intro apply rcases assumption
NewTactic intro apply rcases assumption
-- -- TODO: The new `cases` works differntly. There is also `cases'`
-- example (A B : Prop) : (A ↔ B) → (A → B) := by

@ -50,6 +50,6 @@ Conclusion
"
"
NewTactics apply left right assumption
NewTactic apply left right assumption
NewLemmas not_or_of_imp
NewLemma not_or_of_imp

@ -34,6 +34,6 @@ Conclusion
`rw` hat automatisch `rfl` ausgeführt und dadurch den Beweis beendet.
"
NewTactics rw
NewTactic rw
NewLemmas not_not not_or_of_imp
NewLemma not_not not_or_of_imp

@ -76,6 +76,6 @@ HiddenHint (A : Prop) (B: Prop) (h : ¬ A B) (ha : A) : B =>
HiddenHint (A : Prop) (B: Prop) (ha : A) (ha' : ¬A) : B =>
"Findest du einen Widerspruch?"
NewTactics rfl assumption trivial left right constructor rcases
NewTactic rfl assumption trivial left right constructor rcases
NewLemmas not_not not_or_of_imp
NewLemma not_not not_or_of_imp

@ -24,7 +24,7 @@ Statement
rw [Nat.pow_succ, Nat.mul_succ, add_assoc, h, mul_comm, ←mul_add]
simp
--NewLemmas Nat.pow_succ, Nat.mul_succ, add_assoc, mul_comm, ←mul_add
--NewLemma Nat.pow_succ, Nat.mul_succ, add_assoc, mul_comm, ←mul_add
-- example (n : ) : Even (n^2 + n) := by
-- induction n

@ -34,8 +34,8 @@ Statement Nat.pos_iff_ne_zero
intro
apply Nat.succ_pos
NewTactics simp
NewLemmas Nat.succ_pos
NewTactic simp
NewLemma Nat.succ_pos
Hint : 0 < Nat.zero ↔ Nat.zero ≠ 0 =>
"Den Induktionsanfang kannst du oft mit `simp` lösen."

@ -18,6 +18,6 @@ Statement
(n : ) (h : 2 ≤ n) : n ≠ 0 := by
linarith
NewTactics linarith
NewTactic linarith
NewLemmas Nat.pos_iff_ne_zero
NewLemma Nat.pos_iff_ne_zero

@ -38,4 +38,4 @@ Hint (P : → Prop) (m : ) (hi : P m) : P (Nat.succ m) =>
In diesem Beispiel kannst du `apply` benützen."
NewTactics induction
NewTactic induction

@ -43,9 +43,9 @@ Statement
rw [Fin.sum_univ_castSucc (n := m + 1)]
rfl
OnlyTactics rw rfl
OnlyTactic rw rfl
NewLemmas Fin.sum_univ_castSucc
NewLemma Fin.sum_univ_castSucc
HiddenHint (m : ) :
∑ i : Fin (m + 1), (i : ) + (m + 1) = ∑ i : Fin (Nat.succ m + 1), ↑i =>

@ -49,9 +49,9 @@ Statement
rw [Fin.sum_univ_castSucc (n := m + 1)]
rfl
OnlyTactics rw rfl
OnlyTactic rw rfl
NewLemmas Fin.sum_univ_castSucc
NewLemma Fin.sum_univ_castSucc
Hint (m : ) :
∑ i : Fin (m + 1), (i : ) + (m + 1) = ∑ i : Fin (Nat.succ m + 1), ↑i =>

@ -67,4 +67,4 @@ Hint (A : Prop) (B : Prop) (C : Prop) (h : A ∧ B) : C ∧ A =>
Hint (A : Prop) (B : Prop) (C : Prop) (h : A → C) : C ∧ A =>
"Ein UND im Goal kann mit `constructor` aufgeteilt werden."
NewTactics left right assumption constructor rcases
NewTactic left right assumption constructor rcases

@ -20,4 +20,4 @@ Conclusion
"
"
NewTactics ring
NewTactic ring

@ -20,4 +20,4 @@ Conclusion
"
"
NewTactics ring
NewTactic ring

@ -58,4 +58,4 @@ Hint (n : ) (hn : odd n) (h : ∀ (x : ), (odd x) → even (x + 1)) : even
"`∀ (x : ), (odd x) → even (x + 1)` ist eigentlich das gleiche wie
`(x : ) → `"
NewTactics ring intro unfold
NewTactic ring intro unfold

@ -22,4 +22,4 @@ Statement : ∃ (f : ), ∀ (x : ), f x = 0 := by
Conclusion ""
NewTactics use simp
NewTactic use simp

@ -72,4 +72,4 @@ Statement
simp <;>
ring
NewTactics fin_cases funext
NewTactic fin_cases funext

@ -40,4 +40,4 @@ Statement
-- fin_cases i;
-- simp,
--NewLemmas top_le_span_range_iff_forall_exists_fun le_top
--NewLemma top_le_span_range_iff_forall_exists_fun le_top

@ -39,6 +39,6 @@ Ring zu lösen, funktioniert aber auch auf ``, auch wenn dieses kein Ring ist
(erst `` ist ein Ring).
"
NewTactics ring
NewTactic ring
#eval 4 / 6

@ -67,4 +67,4 @@ Hint (a : ) (b : ) (c : ) (d : ) (h₁ : c = d) (h₂ : d = b) (h₃
Conclusion "Übrigens, mit `rw [h₁, ←h₂]` könnte man mehrere `rw` zusammenfassen."
NewTactics assumption rw
NewTactic assumption rw

@ -37,4 +37,4 @@ Hint (a : ) (b : ) (h : a = b) (g : a + a ^ 2 = b + 1) : a + a ^ 2 = a + 1
Conclusion "Übrigens, mit `rw [h] at *` kann man im weiteren `h` in **allen** Annahmen und
dem Goal umschreiben."
NewTactics assumption rw
NewTactic assumption rw

@ -30,4 +30,4 @@ Hint (y : ) : (2 * y + 1) ^ 2 = 4 * y ^ 2 + 3 * y + 1 + y =>
Conclusion ""
NewTactics ring rw
NewTactic ring rw

@ -28,4 +28,4 @@ Conclusion
aufzurufen, deshalb brauchst du das nur noch selten.
"
NewTactics rfl
NewTactic rfl

@ -67,5 +67,5 @@ Hint (n : ) (x : ) (hx : n = x + x) : n ^ 2 = 2 * x ^ 2 + 2 * x ^ 2 =>
Hint (n : ) (x : ) (hx : n = x + x) : (x + x) ^ 2 = 2 * x ^ 2 + 2 * x ^ 2 =>
"Die Taktik `ring` löst solche Gleichungen."
NewTactics unfold rcases use rw ring
NewDefinitions Even Odd
NewTactic unfold rcases use rw ring
NewDefinition Even Odd

@ -50,4 +50,4 @@ Hint (n : ) (hn : Odd n) (h : ∀ (x : ), (Odd x) → Even (x + 1)) : Even
Conclusion "Für-alle-Statements, Implikationen und Lemmas/Theoreme verhalten sich alle
praktisch gleich. Mehr dazu später."
NewTactics ring intro unfold
NewTactic ring intro unfold

@ -62,6 +62,6 @@ Hint (n : ) : n + n = 2 * n => "Recap: `ring` löst Gleichungen in ``."
Conclusion ""
NewTactics push_neg intro use rw unfold ring
NewDefinitions Even Odd
NewLemmas not_even not_odd not_exists not_forall
NewTactic push_neg intro use rw unfold ring
NewDefinition Even Odd
NewLemma not_even not_odd not_exists not_forall

@ -50,6 +50,6 @@ Statement
Conclusion ""
NewTactics push_neg intro use rw unfold ring
NewDefinitions Even Odd
NewLemmas not_even not_odd not_exists not_forall
NewTactic push_neg intro use rw unfold ring
NewDefinition Even Odd
NewLemma not_even not_odd not_exists not_forall

@ -43,4 +43,4 @@ Statement
rcases h with ⟨_, h₂⟩
assumption
NewLemmas Nat.prime_def_lt''
NewLemma Nat.prime_def_lt''

@ -47,4 +47,4 @@ Aber glaubt bloß nicht, dass Ihr damit auf *diesem* Planeten viel weiterkommt!
Meine Untertanen verstehen `tauto` nicht. Da müsst Ihr Euch schon etwas mehr anstrengen.
"
NewTactics tauto
NewTactic tauto

@ -31,5 +31,5 @@ Conclusion
**Untertan** Ah, richtig. Ja, Sie haben ja so recht. Das vergesse ich immer. Rfl, rfl, rfl …
"
NewTactics rfl
DisabledTactics tauto
NewTactic rfl
DisabledTactic tauto

@ -41,5 +41,5 @@ Conclusion
zerbrochen habe!
"
NewTactics assumption
DisabledTactics tauto
NewTactic assumption
DisabledTactic tauto

@ -34,5 +34,5 @@ Conclusion
**Untertan** Das ging ja schnell. Super! Vielen Dank.
"
NewTactics assumption
DisabledTactics tauto
NewTactic assumption
DisabledTactic tauto

@ -35,5 +35,5 @@ Wie in einer Mathe-Vorlesung?
Das funktioniert nur in einer Handvoll Situationen.
"
NewTactics trivial
DisabledTactics tauto
NewTactic trivial
DisabledTactic tauto

@ -33,5 +33,5 @@ Conclusion
Die Schwester lacht und eilt ihrem Bruder hinterher.
"
NewTactics trivial
DisabledTactics tauto
NewTactic trivial
DisabledTactic tauto

@ -45,5 +45,5 @@ Conclusion
wir haben eine `contradiction` in unserem Annahmen, also folgt jede beliebige Aussage.
"
NewTactics contradiction
DisabledTactics tauto
NewTactic contradiction
DisabledTactic tauto

@ -39,5 +39,5 @@ Und gleich 38, und gleich 39, …
**Du** Ok, ok, verstehe.
"
NewTactics contradiction
DisabledTactics tauto
NewTactic contradiction
DisabledTactic tauto

@ -33,5 +33,5 @@ worin der Widerspruch besteht: Die Annahme `n ≠ 10` ist genau die Negation vo
Man muss `≠` immer als `¬(· = ·)` lesen.
"
NewTactics contradiction
DisabledTactics tauto
NewTactic contradiction
DisabledTactic tauto

@ -54,5 +54,5 @@ Ihm scheinen diese Fragen inzwischen Spaß zu machen.
Oder ist der nur gemalt? Probier mal!
"
NewTactics constructor
DisabledTactics tauto
NewTactic constructor
DisabledTactic tauto

@ -47,5 +47,5 @@ Conclusion
`rcases h with ⟨h₁, ⟨h₂ , h₃⟩⟩`.
"
NewTactics rcases
DisabledTactics tauto
NewTactic rcases
DisabledTactic tauto

@ -44,5 +44,5 @@ Conclusion
Auch dieser Formalosoph zieht zufrieden von dannen.
"
NewTactics left right assumption
DisabledTactics tauto
NewTactic left right assumption
DisabledTactic tauto

@ -125,5 +125,5 @@ Die Worte, die Du aktiv gebrauchen musst, heißen zusammengefasst `Taktiken`. H
"
NewTactics assumption rcases
DisabledTactics tauto
NewTactic assumption rcases
DisabledTactic tauto

@ -74,5 +74,5 @@ Conclusion
Königin *Logisinde* ist in der Zwischenzeit eingeschlafen, und ihr stehlt euch heimlich davon.
"
NewTactics left right assumption constructor rcases rfl contradiction trivial
DisabledTactics tauto
NewTactic left right assumption constructor rcases rfl contradiction trivial
DisabledTactic tauto

@ -39,4 +39,4 @@ Statement mem_univ
trivial
-- tauto
NewTactics tauto trivial
NewTactic tauto trivial

@ -39,7 +39,7 @@ Statement subset_empty_iff
intro h hA
apply mem_univ -- or `trivial`.
NewTactics intro trivial apply
NewTactic intro trivial apply
-- blocked: tauto simp
end MySet

@ -54,8 +54,8 @@ Statement subset_empty_iff
rcases a with ⟨h₁, h₂⟩
assumption
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset
end MySet

@ -45,8 +45,8 @@ Statement eq_empty_iff_forall_not_mem
rw [←subset_empty_iff]
rfl -- This is quite a miracle :)
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset
end MySet

@ -30,6 +30,6 @@ Statement nonempty_iff_ne_empty
push_neg
rfl
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -26,6 +26,6 @@ Statement
"" (A B : Set ) : (A ∅) ∩ B = A ∩ (univ ∩ B) := by
simp
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -30,6 +30,6 @@ Statement
rw [←union_diff_distrib]
rw [univ_union]
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -30,6 +30,6 @@ Statement
rw [←not_mem_compl_iff]
exact h4
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -22,7 +22,7 @@ man die de-Morganschen Regeln einfach selber beweisen könnten.
Die meisten Aufgaben über Mengen sind eine Kombination von `rw` und `simp_rw` verschiedenster
NewLemmas in `import Mathlib.Data.Set`.
NewLemma in `import Mathlib.Data.Set`.
Die Taktik `simp_rw` funktioniert ähnlich wie `rw`, aber sie versucht jedes Lemma so oft
wie möglich anzuwenden. Wir kennen also 4 etwas verwandte Optionen um Lemmas und Theoreme zu
@ -52,7 +52,7 @@ Statement
rw [diff_eq_compl_inter]
rw [inter_comm]
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
-- TODOs
-- Lemmas compl_union compl_inter mem_compl_iff

@ -33,6 +33,6 @@ Statement
rw [Set.mem_diff]
exact hx
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -32,6 +32,6 @@ Statement
({4, 9} : Set ) = Set.insert 4 {9} := by
rfl
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -35,5 +35,5 @@ Statement
simp_rw [mem_insert_iff, mem_singleton_iff] at *
tauto
NewTactics simp_rw intro tauto rw
NewTactic simp_rw intro tauto rw
--Lemmas Subset.antisymm_iff empty_subset

@ -32,6 +32,6 @@ Statement
use 1
ring
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -37,6 +37,6 @@ Statement
assumption
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -36,6 +36,6 @@ Statement
ring
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -34,6 +34,6 @@ Statement
ring
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -28,6 +28,6 @@ Statement
{ x ∈ (S : Set ) | 0 ≤ x} ⊆ S := by
library_search
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -23,6 +23,6 @@ Statement
use 1
ring
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -22,6 +22,6 @@ open Set
Statement
"" : True := sorry
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -23,6 +23,6 @@ Statement
use 1
ring
NewTactics constructor intro rw assumption rcases simp tauto trivial
NewTactic constructor intro rw assumption rcases simp tauto trivial
NewLemmas Subset.antisymm_iff empty_subset
NewLemma Subset.antisymm_iff empty_subset

@ -54,7 +54,7 @@ Statement
rw [←Set.union_diff_distrib]
rw [Set.univ_union]
NewTactics rw
NewTactic rw
example : 4 ∈ (univ : Set ) := by
trivial

@ -35,4 +35,4 @@ lemma asdf (a b c d : ) (h₁ : c = d) (h₂ : a = b) (h₃ : a = d) : b = c
Conclusion ""
NewTactics assumption
NewTactic assumption

@ -44,4 +44,4 @@ Statement
(n : ) : (∑ i : Fin n, (0 + 0)) = 0 := by
simp
NewTactics simp
NewTactic simp

@ -34,7 +34,7 @@ Statement
simp
ring
NewLemmas Finset.sum_add_distrib add_comm
NewLemma Finset.sum_add_distrib add_comm
Hint (n : ) : ∑ x : Fin n, ↑x + ∑ x : Fin n, 1 = n + ∑ i : Fin n, ↑i =>
"Die zweite Summe `∑ x : Fin n, 1` kann `simp` zu `n` vereinfacht werden."

@ -52,8 +52,8 @@ Statement arithmetic_sum
rw [Nat.succ_eq_add_one]
ring
NewTactics induction
NewLemmas Fin.sum_univ_castSucc Nat.succ_eq_add_one mul_add add_mul
NewTactic induction
NewLemma Fin.sum_univ_castSucc Nat.succ_eq_add_one mul_add add_mul
Hint (n : ) : 2 * (∑ i : Fin (n + 1), ↑i) = n * (n + 1) =>
"Als Erinnerung, einen Induktionsbeweis startet man mit `induction n`."

@ -33,4 +33,4 @@ $\\sum_{i=0}^n\\sum_{j=0}^m 2^i (1 + j) = \\sum_{j=0}^m\\sum_{i=0}^n 2^i (1 +
∑ j : Fin m, ∑ i : Fin n, ( 2^i * (1 + j) : ) := by
rw [Finset.sum_comm]
NewLemmas Finset.sum_comm
NewLemma Finset.sum_comm

@ -57,7 +57,7 @@ Statement
rw [arithmetic_sum]
ring
NewLemmas arithmetic_sum add_pow_two
NewLemma arithmetic_sum add_pow_two
HiddenHint (m : ) : ∑ i : Fin (m + 1), (i : ) ^ 3 = (∑ i : Fin (m + 1), ↑i) ^ 2 =>
"Führe auch hier einen Induktionsbeweis."

@ -37,4 +37,4 @@ example (n : ) (h : 5 ≤ n) : n^2 < 2 ^ n
sorry
| n + 5 => by sorry
NewTactics rw simp ring
NewTactic rw simp ring

@ -32,4 +32,4 @@ Statement
-- ring
-- rw [Nat.succ_eq_one_add]
-- rw []
NewTactics rw simp ring
NewTactic rw simp ring

@ -46,4 +46,4 @@ Statement
-- rw [mul_add, hn]
-- ring
NewTactics ring
NewTactic ring

Loading…
Cancel
Save