Theory Real_Vector_Spaces
section ‹Vector Spaces and Algebras over the Reals›
theory Real_Vector_Spaces
imports Real Topological_Spaces Vector_Spaces
begin
subsection ‹Real vector spaces›
class scaleR =
fixes scaleR :: "real ⇒ 'a ⇒ 'a" (infixr ‹*⇩R› 75)
begin
abbreviation divideR :: "'a ⇒ real ⇒ 'a" (infixl ‹'/⇩R› 70)
where "x /⇩R r ≡ inverse r *⇩R x"
end
class real_vector = scaleR + ab_group_add +
assumes scaleR_add_right: "a *⇩R (x + y) = a *⇩R x + a *⇩R y"
and scaleR_add_left: "(a + b) *⇩R x = a *⇩R x + b *⇩R x"
and scaleR_scaleR: "a *⇩R b *⇩R x = (a * b) *⇩R x"
and scaleR_one: "1 *⇩R x = x"
class real_algebra = real_vector + ring +
assumes mult_scaleR_left [simp]: "a *⇩R x * y = a *⇩R (x * y)"
and mult_scaleR_right [simp]: "x * a *⇩R y = a *⇩R (x * y)"
class real_algebra_1 = real_algebra + ring_1
class real_div_algebra = real_algebra_1 + division_ring
class real_field = real_div_algebra + field
instantiation real :: real_field
begin
definition real_scaleR_def [simp]: "scaleR a x = a * x"
instance
by standard (simp_all add: algebra_simps)
end
locale linear = Vector_Spaces.linear "scaleR::_⇒_⇒'a::real_vector" "scaleR::_⇒_⇒'b::real_vector"
begin
lemmas scaleR = scale
end
global_interpretation real_vector?: vector_space "scaleR :: real ⇒ 'a ⇒ 'a :: real_vector"
rewrites "Vector_Spaces.linear (*⇩R) (*⇩R) = linear"
and "Vector_Spaces.linear (*) (*⇩R) = linear"
defines dependent_raw_def: dependent = real_vector.dependent
and representation_raw_def: representation = real_vector.representation
and subspace_raw_def: subspace = real_vector.subspace
and span_raw_def: span = real_vector.span
and extend_basis_raw_def: extend_basis = real_vector.extend_basis
and dim_raw_def: dim = real_vector.dim
proof unfold_locales
show "Vector_Spaces.linear (*⇩R) (*⇩R) = linear" "Vector_Spaces.linear (*) (*⇩R) = linear"
by (force simp: linear_def real_scaleR_def[abs_def])+