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/LinearAlgebra/L08_GeneratingSet.lean

44 lines
1.2 KiB
Plaintext

import Adam.Metadata
2 years ago
import Mathlib.Algebra.Module.Submodule.Lattice
import Mathlib.Data.Real.Basic -- definiert ``
import Mathlib.Data.Fin.VecNotation -- Importiert Matrix/Vektor-Notation
--import Mathlib.LinearAlgebra.FinSupp -- contains `top_le_span_range_iff_forall_exists_fun`
import Mathlib.Tactic.FinCases
import Mathlib.Algebra.BigOperators.Finsupp -- default?
import Mathlib.LinearAlgebra.Span
Game "Adam"
2 years ago
World "Module"
Level 8
Title "Erzeugendensystem"
Introduction
"
Dass eine Familie von Vektoren `M` den ganzen Vektorraum
aufspannt, sagt man in Lean mit ` ≤ span K M`. (Man könnte auch `span K M = ` schreiben,
aber per Konvention wird `≤` bevorzugt, da `x ≤ ` immer gilt (siehe `le_top`))
"
-- notation "ℝ²" => Fin 2 →
open Submodule Set Finsupp
open BigOperators -- Summen Notation
Statement
"Zeige, dass `![1, 0], ![1, 1]` den ganzen ``-Vektorraum `ℝ²` aufspannt."
: ≤ span (Set.range ![(![1, 0] : Fin 2 → ), ![1, 1]]) := by
--rw [top_le_span_range_iff_forall_exists_fun]
sorry
-- intro v,
-- use ![v 0 - v 1, v 1],
-- simp,
-- funext i,
-- fin_cases i;
-- simp,
--NewLemma top_le_span_range_iff_forall_exists_fun le_top