theory Cartesian_Space
imports
Finite_Cartesian_Product Linear_Algebra
begin
definition "cart_basis = {axis i 1 | i. i∈UNIV}"
lemma finite_cart_basis: "finite (cart_basis)" unfolding cart_basis_def
using finite_Atleast_Atmost_nat by fastforce
lemma card_cart_basis: "card (cart_basis::('a::zero_neq_one^'i) set) = CARD('i)"
unfolding cart_basis_def Setcompr_eq_image
by (rule card_image) (auto simp: inj_on_def axis_eq_axis)
interpretation vec: vector_space "( *s) "
by unfold_locales (vector algebra_simps)+
lemma independent_cart_basis:
"vec.independent (cart_basis)"
proof (rule vec.independent_if_scalars_zero)
show "finite (cart_basis)" using finite_cart_basis .
fix f::"('a, 'b) vec ⇒ 'a" and x::"('a, 'b) vec"
assume eq_0: "(∑x∈cart_basis. f x *s x) = 0" and x_in: "x ∈ cart_basis"
obtain i where x: "x = axis i 1" using x_in unfolding cart_basis_def by auto
have sum_eq_0: "(∑x∈(cart_basis) - {x}. f x * (x $ i)) = 0"
proof (rule sum.neutral, rule ballI)
fix xa assume xa: "xa ∈ cart_basis - {x}"
obtain a where a: "xa = axis a 1" and a_not_i: "a ≠ i"
using xa x unfolding cart_basis_def by auto
have "xa $ i = 0" unfolding a axis_def using a_not_i by auto
thus "f xa * xa $ i = 0" by simp
qed
have "0 = (∑x∈cart_basis. f x *s x) $ i" using eq_0 by simp
also have "... = (∑x∈cart_basis. (f x *s x) $ i)" unfolding sum_component ..
also have "... = (∑x∈cart_basis. f x * (x $ i))" unfolding vector_smult_component ..
also have "... = f x * (x $ i) + (∑x∈(cart_basis) - {x}. f x * (x $ i))"
by (rule sum.remove[OF finite_cart_basis x_in])
also have "... = f x * (x $ i)" unfolding sum_eq_0 by simp
also have "... = f x" unfolding x axis_def by auto
finally show "f x = 0" ..
qed
lemma span_cart_basis:
"vec.span (cart_basis) = UNIV"
proof (auto)
fix x::"('a, 'b) vec"
let ?f="λv. x $ (THE i. v = axis i 1)"
show "x ∈ vec.span (cart_basis)"
apply (unfold vec.span_finite[OF finite_cart_basis])
apply (rule image_eqI[of _ _ ?f])
apply (subst vec_eq_iff)
apply clarify
proof -
fix i::'b
let ?w = "axis i (1::'a)"
have the_eq_i: "(THE a. ?w = axis a 1) = i"
by (rule the_equality, auto simp: axis_eq_axis)
have sum_eq_0: "(∑v∈(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i) = 0"
proof (rule sum.neutral, rule ballI)
fix xa::"('a, 'b) vec"
assume xa: "xa ∈ cart_basis - {?w}"
obtain j where j: "xa = axis j 1" and i_not_j: "i ≠ j" using xa unfolding cart_basis_def by auto
have the_eq_j: "(THE i. xa = axis i 1) = j"
proof (rule the_equality)
show "xa = axis j 1" using j .
show "⋀i. xa = axis i 1 ⟹ i = j" by (metis axis_eq_axis j zero_neq_one)
qed
show "x $ (THE i. xa = axis i 1) * xa $ i = 0"
apply (subst (2) j)
unfolding the_eq_j unfolding axis_def using i_not_j by simp
qed
have "(∑v∈cart_basis. x $ (THE i. v = axis i 1) *s v) $ i =
(∑v∈cart_basis. (x $ (THE i. v = axis i 1) *s v) $ i)" unfolding sum_component ..
also have "... = (∑v∈cart_basis. x $ (THE i. v = axis i 1) * v $ i)"
unfolding vector_smult_component ..
also have "... = x $ (THE a. ?w = axis a 1) * ?w $ i + (∑v∈(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i)"
by (rule sum.remove[OF finite_cart_basis], auto simp add: cart_basis_def)
also have "... = x $ (THE a. ?w = axis a 1) * ?w $ i" unfolding sum_eq_0 by simp
also have "... = x $ i" unfolding the_eq_i unfolding axis_def by auto
finally show "x $ i = (∑v∈cart_basis. x $ (THE i. v = axis i 1) *s v) $ i" by simp
qed simp
qed
interpretation vec: finite_dimensional_vector_space "( *s)" "cart_basis"
by (unfold_locales, auto simp add: finite_cart_basis independent_cart_basis span_cart_basis)
lemma matrix_vector_mul_linear_gen[intro, simp]:
"Vector_Spaces.linear ( *s) ( *s) (( *v) A)"
by unfold_locales
(vector matrix_vector_mult_def sum.distrib algebra_simps)+
lemma span_vec_eq: "vec.span X = span X"
and dim_vec_eq: "vec.dim X = dim X"
and dependent_vec_eq: "vec.dependent X = dependent X"
and subspace_vec_eq: "vec.subspace X = subspace X"
for X::"(real^'n) set"
unfolding span_raw_def dim_raw_def dependent_raw_def subspace_raw_def
by (auto simp: scalar_mult_eq_scaleR)
lemma linear_componentwise:
fixes f:: "'a::field ^'m ⇒ 'a ^ 'n"
assumes lf: "Vector_Spaces.linear ( *s) ( *s) f"
shows "(f x)$j = sum (λi. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
proof -
interpret lf: Vector_Spaces.linear "( *s)" "( *s)" f
using lf .
let ?M = "(UNIV :: 'm set)"
let ?N = "(UNIV :: 'n set)"
have fM: "finite ?M" by simp
have "?rhs = (sum (λi. (x$i) *s (f (axis i 1))) ?M)$j"
unfolding sum_component by simp
then show ?thesis
unfolding lf.sum[symmetric] lf.scale[symmetric]
unfolding basis_expansion by auto
qed
interpretation vec: Vector_Spaces.linear "( *s)" "( *s)" "( *v) A"
using matrix_vector_mul_linear_gen.
interpretation vec: finite_dimensional_vector_space_pair "( *s)" cart_basis "( *s)" cart_basis ..
lemma matrix_works:
assumes lf: "Vector_Spaces.linear ( *s) ( *s) f"
shows "matrix f *v x = f (x::'a::field ^ 'n)"
apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute)
apply clarify
apply (rule linear_componentwise[OF lf, symmetric])
done
lemma matrix_of_matrix_vector_mul[simp]: "matrix(λx. A *v (x :: 'a::field ^ 'n)) = A"
by (simp add: matrix_eq matrix_works)
lemma matrix_compose_gen:
assumes lf: "Vector_Spaces.linear ( *s) ( *s) (f::'a::{field}^'n ⇒ 'a^'m)"
and lg: "Vector_Spaces.linear ( *s) ( *s) (g::'a^'m ⇒ 'a^_)"
shows "matrix (g o f) = matrix g ** matrix f"
using lf lg Vector_Spaces.linear_compose[OF lf lg] matrix_works[OF Vector_Spaces.linear_compose[OF lf lg]]
by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
lemma matrix_compose:
assumes "linear (f::real^'n ⇒ real^'m)" "linear (g::real^'m ⇒ real^_)"
shows "matrix (g o f) = matrix g ** matrix f"
using matrix_compose_gen[of f g] assms
by (simp add: linear_def scalar_mult_eq_scaleR)
lemma left_invertible_transpose:
"(∃(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) ⟷ (∃(B). A ** B = mat 1)"
by (metis matrix_transpose_mul transpose_mat transpose_transpose)
lemma right_invertible_transpose:
"(∃(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) ⟷ (∃(B). B ** A = mat 1)"
by (metis matrix_transpose_mul transpose_mat transpose_transpose)
lemma linear_matrix_vector_mul_eq:
"Vector_Spaces.linear ( *s) ( *s) f ⟷ linear (f :: real^'n ⇒ real ^'m)"
by (simp add: scalar_mult_eq_scaleR linear_def)
lemma matrix_vector_mul[simp]:
"Vector_Spaces.linear ( *s) ( *s) g ⟹ (λy. matrix g *v y) = g"
"linear f ⟹ (λx. matrix f *v x) = f"
"bounded_linear f ⟹ (λx. matrix f *v x) = f"
for f :: "real^'n ⇒ real ^'m"
by (simp_all add: ext matrix_works linear_matrix_vector_mul_eq linear_linear)
lemma matrix_left_invertible_injective:
fixes A :: "'a::field^'n^'m"
shows "(∃B. B ** A = mat 1) ⟷ inj (( *v) A)"
proof safe
fix B
assume B: "B ** A = mat 1"
show "inj (( *v) A)"
unfolding inj_on_def
by (metis B matrix_vector_mul_assoc matrix_vector_mul_lid)
next
assume "inj (( *v) A)"
from vec.linear_injective_left_inverse[OF matrix_vector_mul_linear_gen this]
obtain g where "Vector_Spaces.linear ( *s) ( *s) g" and g: "g ∘ ( *v) A = id"
by blast
have "matrix g ** A = mat 1"
by (metis matrix_vector_mul_linear_gen ‹Vector_Spaces.linear ( *s) ( *s) g› g matrix_compose_gen
matrix_eq matrix_id_mat_1 matrix_vector_mul(1))
then show "∃B. B ** A = mat 1"
by metis
qed
lemma matrix_left_invertible_ker:
"(∃B. (B::'a::{field} ^'m^'n) ** (A::'a::{field}^'n^'m) = mat 1) ⟷ (∀x. A *v x = 0 ⟶ x = 0)"
unfolding matrix_left_invertible_injective
using vec.inj_on_iff_eq_0[OF vec.subspace_UNIV, of A]
by (simp add: inj_on_def)
lemma matrix_right_invertible_surjective:
"(∃B. (A::'a::field^'n^'m) ** (B::'a::field^'m^'n) = mat 1) ⟷ surj (λx. A *v x)"
proof -
{ fix B :: "'a ^'m^'n"
assume AB: "A ** B = mat 1"
{ fix x :: "'a ^ 'm"
have "A *v (B *v x) = x"
by (simp add: matrix_vector_mul_assoc AB) }
hence "surj (( *v) A)" unfolding surj_def by metis }
moreover
{ assume sf: "surj (( *v) A)"
from vec.linear_surjective_right_inverse[OF _ this]
obtain g:: "'a ^'m ⇒ 'a ^'n" where g: "Vector_Spaces.linear ( *s) ( *s) g" "( *v) A ∘ g = id"
by blast
have "A ** (matrix g) = mat 1"
unfolding matrix_eq matrix_vector_mul_lid
matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
using g(2) unfolding o_def fun_eq_iff id_def
.
hence "∃B. A ** (B::'a^'m^'n) = mat 1" by blast
}
ultimately show ?thesis unfolding surj_def by blast
qed
lemma matrix_left_invertible_independent_columns:
fixes A :: "'a::{field}^'n^'m"
shows "(∃(B::'a ^'m^'n). B ** A = mat 1) ⟷
(∀c. sum (λi. c i *s column i A) (UNIV :: 'n set) = 0 ⟶ (∀i. c i = 0))"
(is "?lhs ⟷ ?rhs")
proof -
let ?U = "UNIV :: 'n set"
{ assume k: "∀x. A *v x = 0 ⟶ x = 0"
{ fix c i
assume c: "sum (λi. c i *s column i A) ?U = 0" and i: "i ∈ ?U"
let ?x = "χ i. c i"
have th0:"A *v ?x = 0"
using c
by (vector matrix_mult_sum)
from k[rule_format, OF th0] i
have "c i = 0" by (vector vec_eq_iff)}
hence ?rhs by blast }
moreover
{ assume H: ?rhs
{ fix x assume x: "A *v x = 0"
let ?c = "λi. ((x$i ):: 'a)"
from H[rule_format, of ?c, unfolded matrix_mult_sum[symmetric], OF x]
have "x = 0" by vector }
}
ultimately show ?thesis unfolding matrix_left_invertible_ker by auto
qed
lemma matrix_right_invertible_independent_rows:
fixes A :: "'a::{field}^'n^'m"
shows "(∃(B::'a^'m^'n). A ** B = mat 1) ⟷
(∀c. sum (λi. c i *s row i A) (UNIV :: 'm set) = 0 ⟶ (∀i. c i = 0))"
unfolding left_invertible_transpose[symmetric]
matrix_left_invertible_independent_columns
by (simp add:)
lemma matrix_right_invertible_span_columns:
"(∃(B::'a::field ^'n^'m). (A::'a ^'m^'n) ** B = mat 1) ⟷
vec.span (columns A) = UNIV" (is "?lhs = ?rhs")
proof -
let ?U = "UNIV :: 'm set"
have fU: "finite ?U" by simp
have lhseq: "?lhs ⟷ (∀y. ∃(x::'a^'m). sum (λi. (x$i) *s column i A) ?U = y)"
unfolding matrix_right_invertible_surjective matrix_mult_sum surj_def
by (simp add: eq_commute)
have rhseq: "?rhs ⟷ (∀x. x ∈ vec.span (columns A))" by blast
{ assume h: ?lhs
{ fix x:: "'a ^'n"
from h[unfolded lhseq, rule_format, of x] obtain y :: "'a ^'m"
where y: "sum (λi. (y$i) *s column i A) ?U = x" by blast
have "x ∈ vec.span (columns A)"
unfolding y[symmetric] scalar_mult_eq_scaleR
proof (rule vec.span_sum [OF vec.span_scale])
show "column i A ∈ vec.span (columns A)" for i
using columns_def vec.span_superset by auto
qed
}
then have ?rhs unfolding rhseq by blast }
moreover
{ assume h:?rhs
let ?P = "λ(y::'a ^'n). ∃(x::'a^'m). sum (λi. (x$i) *s column i A) ?U = y"
{ fix y
have "y ∈ vec.span (columns A)"
unfolding h by blast
then have "?P y"
proof (induction rule: vec.span_induct_alt)
case base
then show ?case
by (metis (full_types) matrix_mult_sum matrix_vector_mult_0_right)
next
case (step c y1 y2)
from step obtain i where i: "i ∈ ?U" "y1 = column i A"
unfolding columns_def by blast
obtain x:: "'a ^'m" where x: "sum (λi. (x$i) *s column i A) ?U = y2"
using step by blast
let ?x = "(χ j. if j = i then c + (x$i) else (x$j))::'a^'m"
show ?case
proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left if_distribR cong del: if_weak_cong)
fix j
have th: "∀xa ∈ ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))"
using i(1) by (simp add: field_simps)
have "sum (λxa. if xa = i then (c + (x$i)) * ((column xa A)$j)
else (x$xa) * ((column xa A$j))) ?U = sum (λxa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
by (rule sum.cong[OF refl]) (use th in blast)
also have "… = sum (λxa. if xa = i then c * ((column i A)$j) else 0) ?U + sum (λxa. ((x$xa) * ((column xa A)$j))) ?U"
by (simp add: sum.distrib)
also have "… = c * ((column i A)$j) + sum (λxa. ((x$xa) * ((column xa A)$j))) ?U"
unfolding sum.delta[OF fU]
using i(1) by simp
finally show "sum (λxa. if xa = i then (c + (x$i)) * ((column xa A)$j)
else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + sum (λxa. ((x$xa) * ((column xa A)$j))) ?U" .
qed
qed
}
then have ?lhs unfolding lhseq ..
}
ultimately show ?thesis by blast
qed
lemma matrix_left_invertible_span_rows_gen:
"(∃(B::'a^'m^'n). B ** (A::'a::field^'n^'m) = mat 1) ⟷ vec.span (rows A) = UNIV"
unfolding right_invertible_transpose[symmetric]
unfolding columns_transpose[symmetric]
unfolding matrix_right_invertible_span_columns
..
lemma matrix_left_invertible_span_rows:
"(∃(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) ⟷ span (rows A) = UNIV"
using matrix_left_invertible_span_rows_gen[of A] by (simp add: span_vec_eq)
lemma matrix_left_right_inverse:
fixes A A' :: "'a::{field}^'n^'n"
shows "A ** A' = mat 1 ⟷ A' ** A = mat 1"
proof -
{ fix A A' :: "'a ^'n^'n"
assume AA': "A ** A' = mat 1"
have sA: "surj (( *v) A)"
using AA' matrix_right_invertible_surjective by auto
from vec.linear_surjective_isomorphism[OF matrix_vector_mul_linear_gen sA]
obtain f' :: "'a ^'n ⇒ 'a ^'n"
where f': "Vector_Spaces.linear ( *s) ( *s) f'" "∀x. f' (A *v x) = x" "∀x. A *v f' x = x" by blast
have th: "matrix f' ** A = mat 1"
by (simp add: matrix_eq matrix_works[OF f'(1)]
matrix_vector_mul_assoc[symmetric] f'(2)[rule_format])
hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp
hence "matrix f' = A'"
by (simp add: matrix_mul_assoc[symmetric] AA')
hence "matrix f' ** A = A' ** A" by simp
hence "A' ** A = mat 1" by (simp add: th)
}
then show ?thesis by blast
qed
lemma invertible_left_inverse:
fixes A :: "'a::{field}^'n^'n"
shows "invertible A ⟷ (∃(B::'a^'n^'n). B ** A = mat 1)"
by (metis invertible_def matrix_left_right_inverse)
lemma invertible_right_inverse:
fixes A :: "'a::{field}^'n^'n"
shows "invertible A ⟷ (∃(B::'a^'n^'n). A** B = mat 1)"
by (metis invertible_def matrix_left_right_inverse)
lemma invertible_mult:
assumes inv_A: "invertible A"
and inv_B: "invertible B"
shows "invertible (A**B)"
proof -
obtain A' where AA': "A ** A' = mat 1" and A'A: "A' ** A = mat 1"
using inv_A unfolding invertible_def by blast
obtain B' where BB': "B ** B' = mat 1" and B'B: "B' ** B = mat 1"
using inv_B unfolding invertible_def by blast
show ?thesis
proof (unfold invertible_def, rule exI[of _ "B'**A'"], rule conjI)
have "A ** B ** (B' ** A') = A ** (B ** (B' ** A'))"
using matrix_mul_assoc[of A B "(B' ** A')", symmetric] .
also have "... = A ** (B ** B' ** A')" unfolding matrix_mul_assoc[of B "B'" "A'"] ..
also have "... = A ** (mat 1 ** A')" unfolding BB' ..
also have "... = A ** A'" unfolding matrix_mul_lid ..
also have "... = mat 1" unfolding AA' ..
finally show "A ** B ** (B' ** A') = mat (1::'a)" .
have "B' ** A' ** (A ** B) = B' ** (A' ** (A ** B))" using matrix_mul_assoc[of B' A' "(A ** B)", symmetric] .
also have "... = B' ** (A' ** A ** B)" unfolding matrix_mul_assoc[of A' A B] ..
also have "... = B' ** (mat 1 ** B)" unfolding A'A ..
also have "... = B' ** B" unfolding matrix_mul_lid ..
also have "... = mat 1" unfolding B'B ..
finally show "B' ** A' ** (A ** B) = mat 1" .
qed
qed
lemma transpose_invertible:
fixes A :: "real^'n^'n"
assumes "invertible A"
shows "invertible (transpose A)"
by (meson assms invertible_def matrix_left_right_inverse right_invertible_transpose)
lemma vector_matrix_mul_assoc:
fixes v :: "('a::comm_semiring_1)^'n"
shows "(v v* M) v* N = v v* (M ** N)"
proof -
from matrix_vector_mul_assoc
have "transpose N *v (transpose M *v v) = (transpose N ** transpose M) *v v" by fast
thus "(v v* M) v* N = v v* (M ** N)"
by (simp add: matrix_transpose_mul [symmetric])
qed
lemma matrix_scaleR_vector_ac:
fixes A :: "real^('m::finite)^'n"
shows "A *v (k *⇩R v) = k *⇩R A *v v"
by (metis matrix_vector_mult_scaleR transpose_scalar vector_scaleR_matrix_ac vector_transpose_matrix)
lemma scaleR_matrix_vector_assoc:
fixes A :: "real^('m::finite)^'n"
shows "k *⇩R (A *v v) = k *⇩R A *v v"
by (metis matrix_scaleR_vector_ac matrix_vector_mult_scaleR)
locale linear_first_finite_dimensional_vector_space =
l?: Vector_Spaces.linear scaleB scaleC f +
B?: finite_dimensional_vector_space scaleB BasisB
for scaleB :: "('a::field => 'b::ab_group_add => 'b)" (infixr "*b" 75)
and scaleC :: "('a => 'c::ab_group_add => 'c)" (infixr "*c" 75)
and BasisB :: "('b set)"
and f :: "('b=>'c)"
lemma vec_dim_card: "vec.dim (UNIV::('a::{field}^'n) set) = CARD ('n)"
proof -
let ?f="λi::'n. axis i (1::'a)"
have "vec.dim (UNIV::('a::{field}^'n) set) = card (cart_basis::('a^'n) set)"
unfolding vec.dim_UNIV ..
also have "... = card ({i. i∈ UNIV}::('n) set)"
proof (rule bij_betw_same_card[of ?f, symmetric], unfold bij_betw_def, auto)
show "inj (λi::'n. axis i (1::'a))" by (simp add: inj_on_def axis_eq_axis)
fix i::'n
show "axis i 1 ∈ cart_basis" unfolding cart_basis_def by auto
fix x::"'a^'n"
assume "x ∈ cart_basis"
thus "x ∈ range (λi. axis i 1)" unfolding cart_basis_def by auto
qed
also have "... = CARD('n)" by auto
finally show ?thesis .
qed
interpretation vector_space_over_itself: vector_space "( *) :: 'a::field => 'a => 'a"
by unfold_locales (simp_all add: algebra_simps)
lemmas [simp del] = vector_space_over_itself.scale_scale
interpretation vector_space_over_itself: finite_dimensional_vector_space
"( *) :: 'a::field => 'a => 'a" "{1}"
by unfold_locales (auto simp: vector_space_over_itself.span_singleton)
lemma dimension_eq_1[code_unfold]: "vector_space_over_itself.dimension TYPE('a::field)= 1"
unfolding vector_space_over_itself.dimension_def by simp
end