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/Sum/L02_Sum.lean

46 lines
1.2 KiB
Plaintext

This file contains ambiguous Unicode characters!

This file contains ambiguous Unicode characters that may be confused with others in your current locale. If your use case is intentional and legitimate, you can safely ignore this warning. Use the Escape button to highlight these characters.

import TestGame.Metadata
import TestGame.Options.BigOperators
import Mathlib
set_option tactic.hygienic false
Game "TestGame"
World "Sum"
Level 2
Title "endliche Summe"
Introduction
"
Generell sind aber nur solche Lemmas `@[simp]` markiert, klar eine Vereinfachung darstellen.
So ist ein Lemma wie `Finset.sum_add_distrib` kein `simp`-Lemma, da beide Seiten je
nach Situation bevorzugt sein könnte:
$$
\\sum_{i = 0}^n a_i + b_i = \\sum_{i = 0}^n a_i + \\sum_{j = 0}^n b_j
$$
Dieses Lemma kann aber mit `rw` angewendet werden.
"
open BigOperators
Statement
"Zeige dass $\\sum_{i=0}^{n-1} (i + 1) = n + \\sum_{i=0}^{n-1} i$."
(n : ) : ∑ i : Fin n, ((i : ) + 1) = n + (∑ i : Fin n, (i : )) := by
rw [Finset.sum_add_distrib]
simp
ring
NewLemmas 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."
Hint (n : ) : ∑ x : Fin n, ↑x + n = n + ∑ x : Fin n, ↑x =>
"Bis auf Umordnung sind jetzt beide Seiten gleich, darum kann `ring` das Goal schließen.
Alternativ kann man auch mit `rw [add_comm]` dies explizit umordnen."