# HG changeset patch # User immler # Date 1525261778 -7200 # Node ID 493b818e8e1018bbd1fa57ec2bc8fd2ac588bd6e # Parent 0a2a1b6507c17b76786908356c451c7f93ca80a0 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly diff -r 0a2a1b6507c1 -r 493b818e8e10 CONTRIBUTORS --- a/CONTRIBUTORS Wed Apr 18 21:12:50 2018 +0100 +++ b/CONTRIBUTORS Wed May 02 13:49:38 2018 +0200 @@ -6,6 +6,12 @@ Contributions to this Isabelle version -------------------------------------- +* April 2018: Jose Divasón (Universidad de la Rioja), + Jesús Aransay (Universidad de la Rioja), Johannes Hölzl (VU Amsterdam), + Fabian Immler (TUM) + Generalizations in the formalization of linear algebra. + + * March 2018: Florian Haftmann Abstract bit operations push_bit, push_take, push_drop, alongside with an algebraic foundation for bit strings and word types in diff -r 0a2a1b6507c1 -r 493b818e8e10 NEWS --- a/NEWS Wed Apr 18 21:12:50 2018 +0100 +++ b/NEWS Wed May 02 13:49:38 2018 +0200 @@ -257,6 +257,15 @@ * The relator rel_filter on filters has been strengthened to its canonical categorical definition with better properties. INCOMPATIBILITY. +* Generalized linear algebra involving linear, span, dependent, dim +from type class real_vector to locales module and vector_space. +Renamed: + - span_inc ~> span_superset + - span_superset ~> span_base + - span_eq ~> span_eq_iff +INCOMPATIBILITY. + + * HOL-Algebra: renamed (^) to [^] * Session HOL-Analysis: Moebius functions, the Riemann mapping diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Algebra/Congruence.thy --- a/src/HOL/Algebra/Congruence.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Algebra/Congruence.thy Wed May 02 13:49:38 2018 +0200 @@ -4,9 +4,8 @@ *) theory Congruence -imports - Main - "HOL-Library.FuncSet" + imports + Main HOL.FuncSet begin section \Objects\ diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Algebra/Group.thy Wed May 02 13:49:38 2018 +0200 @@ -5,7 +5,7 @@ *) theory Group -imports Complete_Lattice "HOL-Library.FuncSet" +imports Complete_Lattice HOL.FuncSet begin section \Monoids and Groups\ diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Algebra/Order.thy --- a/src/HOL/Algebra/Order.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Algebra/Order.thy Wed May 02 13:49:38 2018 +0200 @@ -7,9 +7,8 @@ *) theory Order -imports - "HOL-Library.FuncSet" - Congruence + imports + Congruence begin section \Orders\ diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Bochner_Integration.thy Wed May 02 13:49:38 2018 +0200 @@ -394,17 +394,19 @@ by (simp add: sum.distrib[symmetric] scaleR_add_right) qed -lemma (in linear) simple_bochner_integral_linear: +lemma simple_bochner_integral_linear: + assumes "linear f" assumes g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f (g x)) = f (simple_bochner_integral M g)" proof - + interpret linear f by fact from g have "simple_bochner_integral M (\x. f (g x)) = (\y\g ` space M. measure M {x \ space M. g x = y} *\<^sub>R f y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2[where p="\x y. f x"] zero elim: simple_bochner_integrable.cases) also have "\ = f (simple_bochner_integral M g)" - by (simp add: simple_bochner_integral_def sum scaleR) + by (simp add: simple_bochner_integral_def sum scale) finally show ?thesis . qed @@ -412,8 +414,7 @@ assumes f: "simple_bochner_integrable M f" shows "simple_bochner_integral M (\x. - f x) = - simple_bochner_integral M f" proof - - interpret linear uminus by unfold_locales auto - from f show ?thesis + from linear_uminus f show ?thesis by (rule simple_bochner_integral_linear) qed @@ -646,7 +647,7 @@ assume "(\i. simple_bochner_integral M (s i)) \ x" with s show "(\i. simple_bochner_integral M (\x. T (s i x))) \ T x" - by (auto intro!: T.tendsto simp: T.simple_bochner_integral_linear) + by (auto intro!: T.tendsto simp: simple_bochner_integral_linear T.linear_axioms) qed lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\x. 0) 0" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Bounded_Linear_Function.thy --- a/src/HOL/Analysis/Bounded_Linear_Function.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Bounded_Linear_Function.thy Wed May 02 13:49:38 2018 +0200 @@ -529,7 +529,7 @@ lemma Blinfun_eq_matrix: "bounded_linear f \ Blinfun f = blinfun_of_matrix (\i j. f j \ i)" by (intro blinfun_euclidean_eqI) (auto simp: blinfun_of_matrix_apply bounded_linear_Blinfun_apply inner_Basis if_distrib - cond_application_beta sum.delta' euclidean_representation + if_distribR sum.delta' euclidean_representation cong: if_cong) text \TODO: generalize (via \compact_cball\)?\ diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Wed May 02 13:49:38 2018 +0200 @@ -1,19 +1,9 @@ section \Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space\ theory Cartesian_Euclidean_Space -imports Finite_Cartesian_Product Derivative +imports Cartesian_Space Derivative begin -lemma norm_le_componentwise: - "(\b. b \ Basis \ abs(x \ b) \ abs(y \ b)) \ norm x \ norm y" - by (auto simp: norm_le euclidean_inner [of x x] euclidean_inner [of y y] abs_le_square_iff power2_eq_square intro!: sum_mono) - -lemma norm_le_componentwise_cart: - fixes x :: "real^'n" - shows "(\i. abs(x$i) \ abs(y$i)) \ norm x \ norm y" - unfolding cart_eq_inner_axis - by (rule norm_le_componentwise) (metis axis_index) - lemma subspace_special_hyperplane: "subspace {x. x $ k = 0}" by (simp add: subspace_def) @@ -31,254 +21,9 @@ qed simp qed simp -subsection\Basic componentwise operations on vectors\ - -instantiation vec :: (times, finite) times -begin - -definition "( * ) \ (\ x y. (\ i. (x$i) * (y$i)))" -instance .. - -end - -instantiation vec :: (one, finite) one -begin - -definition "1 \ (\ i. 1)" -instance .. - -end - -instantiation vec :: (ord, finite) ord -begin - -definition "x \ y \ (\i. x$i \ y$i)" -definition "x < (y::'a^'b) \ x \ y \ \ y \ x" -instance .. - -end - -text\The ordering on one-dimensional vectors is linear.\ - -class cart_one = - assumes UNIV_one: "card (UNIV :: 'a set) = Suc 0" -begin - -subclass finite -proof - from UNIV_one show "finite (UNIV :: 'a set)" - by (auto intro!: card_ge_0_finite) -qed - -end - -instance vec:: (order, finite) order - by standard (auto simp: less_eq_vec_def less_vec_def vec_eq_iff - intro: order.trans order.antisym order.strict_implies_order) - -instance vec :: (linorder, cart_one) linorder -proof - obtain a :: 'b where all: "\P. (\i. P i) \ P a" - proof - - have "card (UNIV :: 'b set) = Suc 0" by (rule UNIV_one) - then obtain b :: 'b where "UNIV = {b}" by (auto iff: card_Suc_eq) - then have "\P. (\i\UNIV. P i) \ P b" by auto - then show thesis by (auto intro: that) - qed - fix x y :: "'a^'b::cart_one" - note [simp] = less_eq_vec_def less_vec_def all vec_eq_iff field_simps - show "x \ y \ y \ x" by auto -qed - -text\Constant Vectors\ - -definition "vec x = (\ i. x)" - lemma interval_cbox_cart: "{a::real^'n..b} = cbox a b" by (auto simp add: less_eq_vec_def mem_box Basis_vec_def inner_axis) -text\Also the scalar-vector multiplication.\ - -definition vector_scalar_mult:: "'a::times \ 'a ^ 'n \ 'a ^ 'n" (infixl "*s" 70) - where "c *s x = (\ i. c * (x$i))" - - -subsection \A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\ - -lemma sum_cong_aux: - "(\x. x \ A \ f x = g x) \ sum f A = sum g A" - by (auto intro: sum.cong) - -hide_fact (open) sum_cong_aux - -method_setup vector = \ -let - val ss1 = - simpset_of (put_simpset HOL_basic_ss @{context} - addsimps [@{thm sum.distrib} RS sym, - @{thm sum_subtractf} RS sym, @{thm sum_distrib_left}, - @{thm sum_distrib_right}, @{thm sum_negf} RS sym]) - val ss2 = - simpset_of (@{context} addsimps - [@{thm plus_vec_def}, @{thm times_vec_def}, - @{thm minus_vec_def}, @{thm uminus_vec_def}, - @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def}, - @{thm scaleR_vec_def}, - @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}]) - fun vector_arith_tac ctxt ths = - simp_tac (put_simpset ss1 ctxt) - THEN' (fn i => resolve_tac ctxt @{thms Cartesian_Euclidean_Space.sum_cong_aux} i - ORELSE resolve_tac ctxt @{thms sum.neutral} i - ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i) - (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) - THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths) -in - Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (vector_arith_tac ctxt ths)) -end -\ "lift trivial vector statements to real arith statements" - -lemma vec_0[simp]: "vec 0 = 0" by vector -lemma vec_1[simp]: "vec 1 = 1" by vector - -lemma vec_inj[simp]: "vec x = vec y \ x = y" by vector - -lemma vec_in_image_vec: "vec x \ (vec ` S) \ x \ S" by auto - -lemma vec_add: "vec(x + y) = vec x + vec y" by vector -lemma vec_sub: "vec(x - y) = vec x - vec y" by vector -lemma vec_cmul: "vec(c * x) = c *s vec x " by vector -lemma vec_neg: "vec(- x) = - vec x " by vector - -lemma vec_scaleR: "vec(c * x) = c *\<^sub>R vec x" - by vector - -lemma vec_sum: - assumes "finite S" - shows "vec(sum f S) = sum (vec \ f) S" - using assms -proof induct - case empty - then show ?case by simp -next - case insert - then show ?case by (auto simp add: vec_add) -qed - -text\Obvious "component-pushing".\ - -lemma vec_component [simp]: "vec x $ i = x" - by vector - -lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i" - by vector - -lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)" - by vector - -lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector - -lemmas vector_component = - vec_component vector_add_component vector_mult_component - vector_smult_component vector_minus_component vector_uminus_component - vector_scaleR_component cond_component - - -subsection \Some frequently useful arithmetic lemmas over vectors\ - -instance vec :: (semigroup_mult, finite) semigroup_mult - by standard (vector mult.assoc) - -instance vec :: (monoid_mult, finite) monoid_mult - by standard vector+ - -instance vec :: (ab_semigroup_mult, finite) ab_semigroup_mult - by standard (vector mult.commute) - -instance vec :: (comm_monoid_mult, finite) comm_monoid_mult - by standard vector - -instance vec :: (semiring, finite) semiring - by standard (vector field_simps)+ - -instance vec :: (semiring_0, finite) semiring_0 - by standard (vector field_simps)+ -instance vec :: (semiring_1, finite) semiring_1 - by standard vector -instance vec :: (comm_semiring, finite) comm_semiring - by standard (vector field_simps)+ - -instance vec :: (comm_semiring_0, finite) comm_semiring_0 .. -instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add .. -instance vec :: (semiring_0_cancel, finite) semiring_0_cancel .. -instance vec :: (comm_semiring_0_cancel, finite) comm_semiring_0_cancel .. -instance vec :: (ring, finite) ring .. -instance vec :: (semiring_1_cancel, finite) semiring_1_cancel .. -instance vec :: (comm_semiring_1, finite) comm_semiring_1 .. - -instance vec :: (ring_1, finite) ring_1 .. - -instance vec :: (real_algebra, finite) real_algebra - by standard (simp_all add: vec_eq_iff) - -instance vec :: (real_algebra_1, finite) real_algebra_1 .. - -lemma of_nat_index: "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" -proof (induct n) - case 0 - then show ?case by vector -next - case Suc - then show ?case by vector -qed - -lemma one_index [simp]: "(1 :: 'a :: one ^ 'n) $ i = 1" - by vector - -lemma neg_one_index [simp]: "(- 1 :: 'a :: {one, uminus} ^ 'n) $ i = - 1" - by vector - -instance vec :: (semiring_char_0, finite) semiring_char_0 -proof - fix m n :: nat - show "inj (of_nat :: nat \ 'a ^ 'b)" - by (auto intro!: injI simp add: vec_eq_iff of_nat_index) -qed - -instance vec :: (numeral, finite) numeral .. -instance vec :: (semiring_numeral, finite) semiring_numeral .. - -lemma numeral_index [simp]: "numeral w $ i = numeral w" - by (induct w) (simp_all only: numeral.simps vector_add_component one_index) - -lemma neg_numeral_index [simp]: "- numeral w $ i = - numeral w" - by (simp only: vector_uminus_component numeral_index) - -instance vec :: (comm_ring_1, finite) comm_ring_1 .. -instance vec :: (ring_char_0, finite) ring_char_0 .. - -lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x" - by (vector mult.assoc) -lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" - by (vector field_simps) -lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" - by (vector field_simps) -lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector -lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector -lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y" - by (vector field_simps) -lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector -lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector -lemma vector_sneg_minus1: "-x = (-1::'a::ring_1) *s x" by vector -lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector -lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x" - by (vector field_simps) - -lemma vec_eq[simp]: "(vec m = vec n) \ (m = n)" - by (simp add: vec_eq_iff) - -lemma linear_vec [simp]: "linear vec" - by (simp add: linearI vec_add vec_eq_iff) - lemma differentiable_vec: fixes S :: "'a::euclidean_space set" shows "vec differentiable_on S" @@ -296,73 +41,6 @@ "box (vec a) (vec b) = {} \ box a b = {}" by (auto simp: Basis_vec_def mem_box box_eq_empty inner_axis) -lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero) - -lemma norm_axis_1 [simp]: "norm (axis m (1::real)) = 1" - by (simp add: inner_axis' norm_eq_1) - -lemma vector_mul_eq_0[simp]: "(a *s x = 0) \ a = (0::'a::idom) \ x = 0" - by vector - -lemma vector_mul_lcancel[simp]: "a *s x = a *s y \ a = (0::real) \ x = y" - by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib) - -lemma vector_mul_rcancel[simp]: "a *s x = b *s x \ (a::real) = b \ x = 0" - by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib) - -lemma vector_mul_lcancel_imp: "a \ (0::real) ==> a *s x = a *s y ==> (x = y)" - by (metis vector_mul_lcancel) - -lemma vector_mul_rcancel_imp: "x \ 0 \ (a::real) *s x = b *s x ==> a = b" - by (metis vector_mul_rcancel) - -lemma component_le_norm_cart: "\x$i\ \ norm x" - apply (simp add: norm_vec_def) - apply (rule member_le_L2_set, simp_all) - done - -lemma norm_bound_component_le_cart: "norm x \ e ==> \x$i\ \ e" - by (metis component_le_norm_cart order_trans) - -lemma norm_bound_component_lt_cart: "norm x < e ==> \x$i\ < e" - by (metis component_le_norm_cart le_less_trans) - -lemma norm_le_l1_cart: "norm x \ sum(\i. \x$i\) UNIV" - by (simp add: norm_vec_def L2_set_le_sum) - -lemma scalar_mult_eq_scaleR [simp]: "c *s x = c *\<^sub>R x" - unfolding scaleR_vec_def vector_scalar_mult_def by simp - -lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \c\ * dist x y" - unfolding dist_norm scalar_mult_eq_scaleR - unfolding scaleR_right_diff_distrib[symmetric] by simp - -lemma sum_component [simp]: - fixes f:: " 'a \ ('b::comm_monoid_add) ^'n" - shows "(sum f S)$i = sum (\x. (f x)$i) S" -proof (cases "finite S") - case True - then show ?thesis by induct simp_all -next - case False - then show ?thesis by simp -qed - -lemma sum_eq: "sum f S = (\ i. sum (\x. (f x)$i ) S)" - by (simp add: vec_eq_iff) - -lemma sum_cmul: - fixes f:: "'c \ ('a::semiring_1)^'n" - shows "sum (\x. c *s f x) S = c *s sum f S" - by (simp add: vec_eq_iff sum_distrib_left) - -lemma sum_norm_allsubsets_bound_cart: - fixes f:: "'a \ real ^'n" - assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (sum f Q) \ e" - shows "sum (\x. norm (f x)) P \ 2 * real CARD('n) * e" - using sum_norm_allsubsets_bound[OF assms] - by simp - subsection\Closures and interiors of halfspaces\ lemma interior_halfspace_le [simp]: @@ -497,270 +175,50 @@ by force qed -subsection \Matrix operations\ - -text\Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"}\ - -definition map_matrix::"('a \ 'b) \ (('a, 'i::finite)vec, 'j::finite) vec \ (('b, 'i)vec, 'j) vec" where - "map_matrix f x = (\ i j. f (x $ i $ j))" - -lemma nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)" - by (simp add: map_matrix_def) - -definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" - (infixl "**" 70) - where "m ** m' == (\ i j. sum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" - -definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" - (infixl "*v" 70) - where "m *v x \ (\ i. sum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" - -definition vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " - (infixl "v*" 70) - where "v v* m == (\ j. sum (\i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" - -definition "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" -definition transpose where - "(transpose::'a^'n^'m \ 'a^'m^'n) A = (\ i j. ((A$j)$i))" -definition "(row::'m => 'a ^'n^'m \ 'a ^'n) i A = (\ j. ((A$i)$j))" -definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\ i. ((A$i)$j))" -definition "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" -definition "columns(A::'a^'n^'m) = { column i A | i. i \ (UNIV :: 'n set)}" - -lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def) -lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)" - by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps) - -lemma matrix_mul_lid [simp]: - fixes A :: "'a::semiring_1 ^ 'm ^ 'n" - shows "mat 1 ** A = A" - apply (simp add: matrix_matrix_mult_def mat_def) - apply vector - apply (auto simp only: if_distrib cond_application_beta sum.delta'[OF finite] - mult_1_left mult_zero_left if_True UNIV_I) - done - -lemma matrix_mul_rid [simp]: - fixes A :: "'a::semiring_1 ^ 'm ^ 'n" - shows "A ** mat 1 = A" - apply (simp add: matrix_matrix_mult_def mat_def) - apply vector - apply (auto simp only: if_distrib cond_application_beta sum.delta[OF finite] - mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) - done - -lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" - apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc) - apply (subst sum.swap) - apply simp - done - -lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" - apply (vector matrix_matrix_mult_def matrix_vector_mult_def - sum_distrib_left sum_distrib_right mult.assoc) - apply (subst sum.swap) - apply simp - done - -lemma matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)" - apply (vector matrix_vector_mult_def mat_def) - apply (simp add: if_distrib cond_application_beta sum.delta' cong del: if_weak_cong) - done - -lemma matrix_transpose_mul: - "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" - by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute) - -lemma matrix_eq: - fixes A B :: "'a::semiring_1 ^ 'n ^ 'm" - shows "A = B \ (\x. A *v x = B *v x)" (is "?lhs \ ?rhs") - apply auto - apply (subst vec_eq_iff) - apply clarify - apply (clarsimp simp add: matrix_vector_mult_def if_distrib cond_application_beta vec_eq_iff cong del: if_weak_cong) - apply (erule_tac x="axis ia 1" in allE) - apply (erule_tac x="i" in allE) - apply (auto simp add: if_distrib cond_application_beta axis_def - sum.delta[OF finite] cong del: if_weak_cong) - done - -lemma matrix_vector_mul_component: "((A::real^_^_) *v x)$k = (A$k) \ x" - by (simp add: matrix_vector_mult_def inner_vec_def) - -lemma dot_lmul_matrix: "((x::real ^_) v* A) \ y = x \ (A *v y)" - apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def sum_distrib_right sum_distrib_left ac_simps) - apply (subst sum.swap) - apply simp - done - -lemma transpose_mat [simp]: "transpose (mat n) = mat n" - by (vector transpose_def mat_def) - -lemma transpose_transpose [simp]: "transpose(transpose A) = A" - by (vector transpose_def) - -lemma row_transpose [simp]: - fixes A:: "'a::semiring_1^_^_" - shows "row i (transpose A) = column i A" - by (simp add: row_def column_def transpose_def vec_eq_iff) - -lemma column_transpose [simp]: - fixes A:: "'a::semiring_1^_^_" - shows "column i (transpose A) = row i A" - by (simp add: row_def column_def transpose_def vec_eq_iff) - -lemma rows_transpose [simp]: "rows(transpose (A::'a::semiring_1^_^_)) = columns A" - by (auto simp add: rows_def columns_def row_transpose intro: set_eqI) - -lemma columns_transpose [simp]: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" - by (metis transpose_transpose rows_transpose) - lemma matrix_mult_transpose_dot_column: fixes A :: "real^'n^'n" - shows "transpose A ** A = (\ i j. (column i A) \ (column j A))" + shows "transpose A ** A = (\ i j. inner (column i A) (column j A))" by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def column_def inner_vec_def) lemma matrix_mult_transpose_dot_row: fixes A :: "real^'n^'n" - shows "A ** transpose A = (\ i j. (row i A) \ (row j A))" + shows "A ** transpose A = (\ i j. inner (row i A) (row j A))" by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def row_def inner_vec_def) text\Two sometimes fruitful ways of looking at matrix-vector multiplication.\ -lemma matrix_mult_dot: "A *v x = (\ i. A$i \ x)" +lemma matrix_mult_dot: "A *v x = (\ i. inner (A$i) x)" by (simp add: matrix_vector_mult_def inner_vec_def) -lemma matrix_mult_sum: - "(A::'a::comm_semiring_1^'n^'m) *v x = sum (\i. (x$i) *s column i A) (UNIV:: 'n set)" - by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute) - -lemma vector_componentwise: - "(x::'a::ring_1^'n) = (\ j. \i\UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)" - by (simp add: axis_def if_distrib sum.If_cases vec_eq_iff) - -lemma basis_expansion: "sum (\i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)" - by (auto simp add: axis_def vec_eq_iff if_distrib sum.If_cases cong del: if_weak_cong) - -lemma linear_componentwise_expansion: - fixes f:: "real ^'m \ real ^ _" - assumes lf: "linear f" - shows "(f x)$j = sum (\i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") -proof - - let ?M = "(UNIV :: 'm set)" - let ?N = "(UNIV :: 'n set)" - have "?rhs = (sum (\i.(x$i) *\<^sub>R f (axis i 1) ) ?M)$j" - unfolding sum_component by simp - then show ?thesis - unfolding linear_sum_mul[OF lf, symmetric] - unfolding scalar_mult_eq_scaleR[symmetric] - unfolding basis_expansion - by simp -qed - -subsection\Inverse matrices (not necessarily square)\ - -definition - "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" - -definition - "matrix_inv(A:: 'a::semiring_1^'n^'m) = - (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" - -text\Correspondence between matrices and linear operators.\ - -definition matrix :: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" - where "matrix f = (\ i j. (f(axis j 1))$i)" - -lemma matrix_id_mat_1: "matrix id = mat 1" - by (simp add: mat_def matrix_def axis_def) - -lemma matrix_scaleR: "(matrix (( *\<^sub>R) r)) = mat r" - by (simp add: mat_def matrix_def axis_def if_distrib cong: if_cong) - -lemma matrix_vector_mul_linear: "linear(\x. A *v (x::real ^ _))" - by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff - field_simps sum_distrib_left sum.distrib) - -lemma - fixes A :: "real^'n^'m" - shows matrix_vector_mult_linear_continuous_at [continuous_intros]: "isCont (( *v) A) z" - and matrix_vector_mult_linear_continuous_on [continuous_intros]: "continuous_on S (( *v) A)" - by (simp_all add: linear_linear linear_continuous_at linear_continuous_on matrix_vector_mul_linear) - -lemma matrix_vector_mult_add_distrib [algebra_simps]: - "A *v (x + y) = A *v x + A *v y" - by (vector matrix_vector_mult_def sum.distrib distrib_left) - -lemma matrix_vector_mult_diff_distrib [algebra_simps]: - fixes A :: "'a::ring_1^'n^'m" - shows "A *v (x - y) = A *v x - A *v y" - by (vector matrix_vector_mult_def sum_subtractf right_diff_distrib) - -lemma matrix_vector_mult_scaleR[algebra_simps]: - fixes A :: "real^'n^'m" - shows "A *v (c *\<^sub>R x) = c *\<^sub>R (A *v x)" - using linear_iff matrix_vector_mul_linear by blast - -lemma matrix_vector_mult_0_right [simp]: "A *v 0 = 0" - by (simp add: matrix_vector_mult_def vec_eq_iff) - -lemma matrix_vector_mult_0 [simp]: "0 *v w = 0" - by (simp add: matrix_vector_mult_def vec_eq_iff) - -lemma matrix_vector_mult_add_rdistrib [algebra_simps]: - "(A + B) *v x = (A *v x) + (B *v x)" - by (vector matrix_vector_mult_def sum.distrib distrib_right) - -lemma matrix_vector_mult_diff_rdistrib [algebra_simps]: - fixes A :: "'a :: ring_1^'n^'m" - shows "(A - B) *v x = (A *v x) - (B *v x)" - by (vector matrix_vector_mult_def sum_subtractf left_diff_distrib) - -lemma matrix_works: - assumes lf: "linear f" - shows "matrix f *v x = f (x::real ^ 'n)" - apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute) - by (simp add: linear_componentwise_expansion lf) - -lemma matrix_vector_mul: "linear f ==> f = (\x. matrix f *v (x::real ^ 'n))" - by (simp add: ext matrix_works) - -declare matrix_vector_mul [symmetric, simp] - -lemma matrix_of_matrix_vector_mul [simp]: "matrix(\x. A *v (x :: real ^ 'n)) = A" - by (simp add: matrix_eq matrix_vector_mul_linear matrix_works) - -lemma matrix_compose: - assumes lf: "linear (f::real^'n \ real^'m)" - and lg: "linear (g::real^'m \ real^_)" - shows "matrix (g \ f) = matrix g ** matrix f" - using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]] - by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def) - -lemma matrix_vector_column: - "(A::'a::comm_semiring_1^'n^_) *v x = sum (\i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" - by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute) - lemma adjoint_matrix: "adjoint(\x. (A::real^'n^'m) *v x) = (\x. transpose A *v x)" apply (rule adjoint_unique) apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def sum_distrib_right sum_distrib_left) apply (subst sum.swap) - apply (auto simp add: ac_simps) + apply (simp add: ac_simps) done lemma matrix_adjoint: assumes lf: "linear (f :: real^'n \ real ^'m)" shows "matrix(adjoint f) = transpose(matrix f)" - apply (subst matrix_vector_mul[OF lf]) - unfolding adjoint_matrix matrix_of_matrix_vector_mul - apply rule - done +proof - + have "matrix(adjoint f) = matrix(adjoint (( *v) (matrix f)))" + by (simp add: lf) + also have "\ = transpose(matrix f)" + unfolding adjoint_matrix matrix_of_matrix_vector_mul + apply rule + done + finally show ?thesis . +qed -lemma inj_matrix_vector_mult: - fixes A::"'a::field^'n^'m" - assumes "invertible A" - shows "inj (( *v) A)" - by (metis assms inj_on_inverseI invertible_def matrix_vector_mul_assoc matrix_vector_mul_lid) +lemma matrix_vector_mul_bounded_linear[intro, simp]: "bounded_linear (( *v) A)" for A :: "real^'n^'m" + using matrix_vector_mul_linear[of A] + by (simp add: linear_conv_bounded_linear linear_matrix_vector_mul_eq) + +lemma + fixes A :: "real^'n^'m" + shows matrix_vector_mult_linear_continuous_at [continuous_intros]: "isCont (( *v) A) z" + and matrix_vector_mult_linear_continuous_on [continuous_intros]: "continuous_on S (( *v) A)" + by (simp_all add: linear_continuous_at linear_continuous_on) subsection\Some bounds on components etc. relative to operator norm\ @@ -769,12 +227,10 @@ fixes A :: "real^'n^'m" shows "norm(column i A) \ onorm(( *v) A)" proof - - have bl: "bounded_linear (( *v) A)" - by (simp add: linear_linear matrix_vector_mul_linear) have "norm (\ j. A $ j $ i) \ norm (A *v axis i 1)" by (simp add: matrix_mult_dot cart_eq_inner_axis) also have "\ \ onorm (( *v) A)" - using onorm [OF bl, of "axis i 1"] by auto + using onorm [OF matrix_vector_mul_bounded_linear, of A "axis i 1"] by auto finally have "norm (\ j. A $ j $ i) \ onorm (( *v) A)" . then show ?thesis unfolding column_def . @@ -794,7 +250,7 @@ lemma component_le_onorm: fixes f :: "real^'m \ real^'n" shows "linear f \ \matrix f $ i $ j\ \ onorm f" - by (metis matrix_component_le_onorm matrix_vector_mul) + by (metis linear_matrix_vector_mul_eq matrix_component_le_onorm matrix_vector_mul) lemma onorm_le_matrix_component_sum: fixes A :: "real^'n^'m" @@ -892,16 +348,8 @@ lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \ (x - ((b \ x) / (b \ b)) *s b) = 0" unfolding inner_simps scalar_mult_eq_scaleR by auto -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 matrix_left_invertible_injective: - fixes A :: "real^'n^'m" + fixes A :: "'a::field^'n^'m" shows "(\B. B ** A = mat 1) \ inj (( *v) A)" proof safe fix B @@ -911,34 +359,29 @@ by (metis B matrix_vector_mul_assoc matrix_vector_mul_lid) next assume "inj (( *v) A)" - with linear_injective_left_inverse[OF matrix_vector_mul_linear] - obtain g where "linear g" and g: "g \ ( *v) A = id" + 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 \linear g\ g matrix_compose matrix_id_mat_1 matrix_of_matrix_vector_mul matrix_vector_mul_linear) + 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::real ^'m^'n) ** (A::real^'n^'m) = mat 1) \ (\x. A *v x = 0 \ x = 0)" - unfolding matrix_left_invertible_injective - using linear_injective_0[OF matrix_vector_mul_linear, of A] - by (simp add: inj_on_def) - lemma matrix_right_invertible_surjective: - "(\B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \ surj (\x. A *v x)" + "(\B. (A::'a::field^'n^'m) ** (B::'a::field^'m^'n) = mat 1) \ surj (\x. A *v x)" proof - - { fix B :: "real ^'m^'n" + { fix B :: "'a ^'m^'n" assume AB: "A ** B = mat 1" - { fix x :: "real ^ 'm" + { fix x :: "'a ^ 'm" have "A *v (B *v x) = x" - by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB) } + by (simp add: matrix_vector_mul_assoc AB) } hence "surj (( *v) A)" unfolding surj_def by metis } moreover { assume sf: "surj (( *v) A)" - from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf] - obtain g:: "real ^'m \ real ^'n" where g: "linear g" "( *v) A \ g = id" + 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" @@ -946,69 +389,32 @@ 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::real^'m^'n) = mat 1" by blast + 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 :: "real^'n^'m" - shows "(\(B::real ^'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 - unfolding matrix_mult_sum vec_eq_iff - by auto - 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 ):: real)" - 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 blast -qed - -lemma matrix_right_invertible_independent_rows: - fixes A :: "real^'n^'m" - shows "(\(B::real^'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: column_transpose) - lemma matrix_right_invertible_span_columns: - "(\(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \ - span (columns A) = UNIV" (is "?lhs = ?rhs") + "(\(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::real^'m). sum (\i. (x$i) *s column i A) ?U = y)" + 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 apply (subst eq_commute) apply rule done - have rhseq: "?rhs \ (\x. x \ span (columns A))" by blast + have rhseq: "?rhs \ (\x. x \ vec.span (columns A))" by blast { assume h: ?lhs - { fix x:: "real ^'n" - from h[unfolded lhseq, rule_format, of x] obtain y :: "real ^'m" + { 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 \ span (columns A)" + have "x \ vec.span (columns A)" unfolding y[symmetric] - apply (rule span_sum) - unfolding scalar_mult_eq_scaleR - apply (rule span_mul) - apply (rule span_superset) + apply (rule vec.span_sum) + apply (rule vec.span_scale) + apply (rule vec.span_base) unfolding columns_def apply blast done @@ -1016,22 +422,24 @@ then have ?rhs unfolding rhseq by blast } moreover { assume h:?rhs - let ?P = "\(y::real ^'n). \(x::real^'m). sum (\i. (x$i) *s column i A) ?U = y" + let ?P = "\(y::'a ^'n). \(x::'a^'m). sum (\i. (x$i) *s column i A) ?U = y" { fix y - have "?P y" - proof (rule span_induct_alt[of ?P "columns A", folded scalar_mult_eq_scaleR]) - show "\x::real ^ 'm. sum (\i. (x$i) *s column i A) ?U = 0" + have "y \ vec.span (columns A)" + unfolding h by blast + then have "?P y" + proof (induction rule: vec.span_induct_alt) + show "\x::'a ^ 'm. sum (\i. (x$i) *s column i A) ?U = 0" by (rule exI[where x=0], simp) next fix c y1 y2 assume y1: "y1 \ columns A" and y2: "?P y2" from y1 obtain i where i: "i \ ?U" "y1 = column i A" unfolding columns_def by blast - from y2 obtain x:: "real ^'m" where + from y2 obtain x:: "'a ^'m" where x: "sum (\i. (x$i) *s column i A) ?U = y2" by blast - let ?x = "(\ j. if j = i then c + (x$i) else (x$j))::real^'m" + let ?x = "(\ j. if j = i then c + (x$i) else (x$j))::'a^'m" show "?P (c*s y1 + y2)" - proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left cond_application_beta cong del: if_weak_cong) + 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))" @@ -1049,9 +457,6 @@ 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 - next - show "y \ span (columns A)" - unfolding h by blast qed } then have ?lhs unfolding lhseq .. @@ -1059,41 +464,20 @@ ultimately show ?thesis by blast qed -lemma matrix_left_invertible_span_rows: - "(\(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \ span (rows A) = UNIV" +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) + + text \The same result in terms of square matrices.\ -lemma matrix_left_right_inverse: - fixes A A' :: "real ^'n^'n" - shows "A ** A' = mat 1 \ A' ** A = mat 1" -proof - - { fix A A' :: "real ^'n^'n" - assume AA': "A ** A' = mat 1" - have sA: "surj (( *v) A)" - unfolding surj_def - apply clarify - apply (rule_tac x="(A' *v y)" in exI) - apply (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid) - done - from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA] - obtain f' :: "real ^'n \ real ^'n" - where f': "linear 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] matrix_vector_mul_lid 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' matrix_mul_rid matrix_mul_lid) - hence "matrix f' ** A = A' ** A" by simp - hence "A' ** A = mat 1" by (simp add: th) - } - then show ?thesis by blast -qed text \Considering an n-element vector as an n-by-1 or 1-by-n matrix.\ @@ -1338,8 +722,8 @@ using connected_ivt_hyperplane[of s x y "axis k 1" a] by (auto simp add: inner_axis inner_commute) -lemma subspace_substandard_cart: "subspace {x::real^_. (\i. P i \ x$i = 0)}" - unfolding subspace_def by auto +lemma subspace_substandard_cart: "vec.subspace {x. (\i. P i \ x$i = 0)}" + unfolding vec.subspace_def by auto lemma closed_substandard_cart: "closed {x::'a::real_normed_vector ^ 'n. \i. P i \ x$i = 0}" @@ -1351,32 +735,49 @@ unfolding Collect_all_eq by (simp add: closed_INT) qed -lemma dim_substandard_cart: "dim {x::real^'n. \i. i \ d \ x$i = 0} = card d" - (is "dim ?A = _") -proof - - let ?a = "\x. axis x 1 :: real^'n" - have *: "{x. \i\Basis. i \ ?a ` d \ x \ i = 0} = ?A" - by (auto simp: image_iff Basis_vec_def axis_eq_axis inner_axis) - have "?a ` d \ Basis" - by (auto simp: Basis_vec_def) - thus ?thesis - using dim_substandard[of "?a ` d"] card_image[of ?a d] - by (auto simp: axis_eq_axis inj_on_def *) -qed +lemma dim_substandard_cart: "vec.dim {x::'a::field^'n. \i. i \ d \ x$i = 0} = card d" + (is "vec.dim ?A = _") +proof (rule vec.dim_unique) + let ?B = "((\x. axis x 1) ` d)" + have subset_basis: "?B \ cart_basis" + by (auto simp: cart_basis_def) + show "?B \ ?A" + by (auto simp: axis_def) + show "vec.independent ((\x. axis x 1) ` d)" + using subset_basis + by (rule vec.independent_mono[OF vec.independent_Basis]) + have "x \ vec.span ?B" if "\i. i \ d \ x $ i = 0" for x::"'a^'n" + proof - + have "finite ?B" + using subset_basis finite_cart_basis + by (rule finite_subset) + have "x = (\i\UNIV. x $ i *s axis i 1)" + by (rule basis_expansion[symmetric]) + also have "\ = (\i\d. (x $ i) *s axis i 1)" + by (rule sum.mono_neutral_cong_right) (auto simp: that) + also have "\ \ vec.span ?B" + by (simp add: vec.span_sum vec.span_clauses) + finally show "x \ vec.span ?B" . + qed + then show "?A \ vec.span ?B" by auto +qed (simp add: card_image inj_on_def axis_eq_axis) + +lemma dim_subset_UNIV_cart_gen: + fixes S :: "('a::field^'n) set" + shows "vec.dim S \ CARD('n)" + by (metis vec.dim_eq_full vec.dim_subset_UNIV vec.span_UNIV vec_dim_card) lemma dim_subset_UNIV_cart: fixes S :: "(real^'n) set" shows "dim S \ CARD('n)" - by (metis dim_subset_UNIV DIM_cart DIM_real mult.right_neutral) + using dim_subset_UNIV_cart_gen[of S] by (simp add: dim_vec_eq) lemma affinity_inverses: assumes m0: "m \ (0::'a::field)" shows "(\x. m *s x + c) \ (\x. inverse(m) *s x + (-(inverse(m) *s c))) = id" "(\x. inverse(m) *s x + (-(inverse(m) *s c))) \ (\x. m *s x + c) = id" using m0 - apply (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff) - apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1 [symmetric]) - done + by (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff) lemma vector_affinity_eq: assumes m0: "(m::'a::field) \ 0" @@ -1580,13 +981,18 @@ text\Equivalence of row and column rank is taken from George Mackiw's paper, Mathematics Magazine 1995, p. 285.\ +lemma matrix_vector_mult_in_columnspace_gen: + fixes A :: "'a::field^'n^'m" + shows "(A *v x) \ vec.span(columns A)" + apply (simp add: matrix_vector_column columns_def transpose_def column_def) + apply (intro vec.span_sum vec.span_scale) + apply (force intro: vec.span_base) + done + lemma matrix_vector_mult_in_columnspace: fixes A :: "real^'n^'m" shows "(A *v x) \ span(columns A)" - apply (simp add: matrix_vector_column columns_def transpose_def column_def) - apply (intro span_sum span_mul) - apply (force intro: span_superset) - done + using matrix_vector_mult_in_columnspace_gen[of A x] by (simp add: span_vec_eq) lemma orthogonal_nullspace_rowspace: fixes A :: "real^'n^'m" @@ -1608,16 +1014,20 @@ lemma nullspace_inter_rowspace: fixes A :: "real^'n^'m" shows "A *v x = 0 \ x \ span(rows A) \ x = 0" - using orthogonal_nullspace_rowspace orthogonal_self by auto + using orthogonal_nullspace_rowspace orthogonal_self span_zero matrix_vector_mult_0_right + by blast lemma matrix_vector_mul_injective_on_rowspace: fixes A :: "real^'n^'m" shows "\A *v x = A *v y; x \ span(rows A); y \ span(rows A)\ \ x = y" using nullspace_inter_rowspace [of A "x-y"] - by (metis eq_iff_diff_eq_0 matrix_vector_mult_diff_distrib span_diff) + by (metis diff_eq_diff_eq diff_self matrix_vector_mult_diff_distrib span_diff) -definition rank :: "real^'n^'m=>nat" - where "rank A \ dim(columns A)" +definition rank :: "'a::field^'n^'m=>nat" + where row_rank_def_gen: "rank A \ vec.dim(rows A)" + +lemma row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m" + by (auto simp: row_rank_def_gen dim_vec_eq) lemma dim_rows_le_dim_columns: fixes A :: "real^'n^'m" @@ -1628,31 +1038,35 @@ obtain B where "independent B" "span(rows A) \ span B" and B: "B \ span(rows A)""card B = dim (span(rows A))" using basis_exists [of "span(rows A)"] by blast - with span_subspace have eq: "span B = span(rows A)" - by auto + then have eq: "span B = span(rows A)" + using span_subspace subspace_span by blast then have inj: "inj_on (( *v) A) (span B)" - using inj_on_def matrix_vector_mul_injective_on_rowspace by blast + by (simp add: inj_on_def matrix_vector_mul_injective_on_rowspace) then have ind: "independent (( *v) A ` B)" - by (rule independent_inj_on_image [OF \independent B\ matrix_vector_mul_linear]) - then have "finite (( *v) A ` B) \ card (( *v) A ` B) \ dim (( *v) A ` B)" - by (rule independent_bound_general) - then show ?thesis - by (metis (no_types, lifting) B ind inj eq card_image image_subset_iff independent_card_le_dim inj_on_subset matrix_vector_mult_in_columnspace) + by (rule linear_independent_injective_image [OF Finite_Cartesian_Product.matrix_vector_mul_linear \independent B\]) + have "dim (span (rows A)) \ card (( *v) A ` B)" + unfolding B(2)[symmetric] + using inj + by (auto simp: card_image inj_on_subset span_superset) + also have "\ \ dim (span (columns A))" + using _ ind + by (rule independent_card_le_dim) (auto intro!: matrix_vector_mult_in_columnspace) + finally show ?thesis . qed then show ?thesis - by simp + by (simp add: dim_span) qed -lemma rank_row: +lemma column_rank_def: fixes A :: "real^'n^'m" - shows "rank A = dim(rows A)" - unfolding rank_def - by (metis dim_rows_le_dim_columns columns_transpose dual_order.antisym rows_transpose) + shows "rank A = dim(columns A)" + unfolding row_rank_def + by (metis columns_transpose dim_rows_le_dim_columns le_antisym rows_transpose) lemma rank_transpose: fixes A :: "real^'n^'m" shows "rank(transpose A) = rank A" - by (metis rank_def rank_row rows_transpose) + by (metis column_rank_def row_rank_def rows_transpose) lemma matrix_vector_mult_basis: fixes A :: "real^'n^'m" @@ -1667,22 +1081,26 @@ lemma rank_dim_range: fixes A :: "real^'n^'m" shows "rank A = dim(range (\x. A *v x))" - unfolding rank_def + unfolding column_rank_def proof (rule span_eq_dim) - show "span (columns A) = span (range (( *v) A))" - apply (auto simp: columns_image_basis span_linear_image matrix_vector_mul_linear) - by (metis columns_image_basis matrix_vector_mul_linear matrix_vector_mult_in_columnspace span_linear_image) + have "span (columns A) \ span (range (( *v) A))" (is "?l \ ?r") + by (simp add: columns_image_basis image_subsetI span_mono) + then show "?l = ?r" + by (metis (no_types, lifting) image_subset_iff matrix_vector_mult_in_columnspace + span_eq span_span) qed lemma rank_bound: fixes A :: "real^'n^'m" shows "rank A \ min CARD('m) (CARD('n))" - by (metis (mono_tags, hide_lams) min.bounded_iff DIM_cart DIM_real dim_subset_UNIV mult.right_neutral rank_def rank_transpose) + by (metis (mono_tags, lifting) dim_subset_UNIV_cart min.bounded_iff + column_rank_def row_rank_def) lemma full_rank_injective: fixes A :: "real^'n^'m" shows "rank A = CARD('n) \ inj (( *v) A)" - by (simp add: matrix_left_invertible_injective [symmetric] matrix_left_invertible_span_rows rank_row dim_eq_full [symmetric]) + by (simp add: matrix_left_invertible_injective [symmetric] matrix_left_invertible_span_rows row_rank_def + dim_eq_full [symmetric] card_cart_basis vec.dimension_def) lemma full_rank_surjective: fixes A :: "real^'n^'m" @@ -1703,10 +1121,10 @@ shows "(\x. (x \ 0) \ A *v x = 0) \ ~(rank A = CARD('n))" by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker) -lemma rank_eq_0: "rank A = 0 \ A = 0" and rank_0 [simp]: "rank 0 = 0" +lemma rank_eq_0: "rank A = 0 \ A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0" + for A :: "real^'n^'m" by (auto simp: rank_dim_range matrix_eq) - lemma rank_mul_le_right: fixes A :: "real^'n^'m" and B :: "real^'p^'n" shows "rank(A ** B) \ rank B" @@ -1714,7 +1132,7 @@ have "rank(A ** B) \ dim (( *v) A ` range (( *v) B))" by (auto simp: rank_dim_range image_comp o_def matrix_vector_mul_assoc) also have "\ \ rank B" - by (simp add: rank_dim_range matrix_vector_mul_linear dim_image_le) + by (simp add: rank_dim_range dim_image_le) finally show ?thesis . qed @@ -1845,4 +1263,4 @@ bounded_linear.uniform_limit[OF bounded_linear_vec_nth] bounded_linear.uniform_limit[OF bounded_linear_component_cart] -end +end \ No newline at end of file diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Cartesian_Space.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Cartesian_Space.thy Wed May 02 13:49:38 2018 +0200 @@ -0,0 +1,312 @@ +(* Title: Cartesian_Space.thy + Author: Amine Chaieb, University of Cambridge + Author: Jose Divasón + Author: Jesús Aransay + Author: Johannes Hölzl, VU Amsterdam + Author: Fabian Immler, TUM +*) +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 + +(*Some interpretations:*) +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 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 matrix_left_invertible_injective: + "(\B. (B::'a::field^'m^'n) ** (A::'a::field^'n^'m) = mat 1) + \ (\x y. A *v x = A *v y \ x = y)" +proof - + { fix B:: "'a^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y" + from xy have "B*v (A *v x) = B *v (A*v y)" by simp + hence "x = y" + unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid . } + moreover + { assume A: "\x y. A *v x = A *v y \ x = y" + hence i: "inj (( *v) A)" unfolding inj_on_def by auto + from vec.linear_exists_left_inverse_on[OF matrix_vector_mul_linear_gen vec.subspace_UNIV i] + obtain g where g: "Vector_Spaces.linear ( *s) ( *s) g" "g o (( *v) A) = id" by (auto simp: id_def module_hom_iff_linear o_def) + have "matrix g ** A = mat 1" + unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] + using g(2) by (metis comp_apply id_apply) + then have "\B. (B::'a::{field}^'m^'n) ** A = mat 1" by blast } + ultimately show ?thesis by blast +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_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 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 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_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)" + unfolding surj_def + apply clarify + apply (rule_tac x="(A' *v y)" in exI) + apply (simp add: matrix_vector_mul_assoc AA') + done + 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) + +(*Finally, some interesting theorems and interpretations that don't appear in any file of the + library.*) + +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 + +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 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) + +end \ No newline at end of file diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Change_Of_Vars.thy Wed May 02 13:49:38 2018 +0200 @@ -179,6 +179,10 @@ by (rule induct_matrix_elementary) (auto intro: assms *) qed +lemma matrix_vector_mult_matrix_matrix_mult_compose: + "( *v) (A ** B) = ( *v) A \ ( *v) B" + by (auto simp: matrix_vector_mul_assoc) + lemma induct_linear_elementary: fixes f :: "real^'n \ real^'n" assumes "linear f" @@ -194,12 +198,15 @@ fix A B assume "P (( *v) A)" and "P (( *v) B)" then show "P (( *v) (A ** B))" - by (metis (no_types, lifting) comp linear_compose matrix_compose matrix_eq matrix_vector_mul matrix_vector_mul_linear) + by (auto simp add: matrix_vector_mult_matrix_matrix_mult_compose matrix_vector_mul_linear + intro!: comp) next fix A :: "real^'n^'n" and i assume "row i A = 0" - then show "P (( *v) A)" - by (metis inner_zero_left matrix_vector_mul_component matrix_vector_mul_linear row_def vec_eq_iff vec_lambda_beta zeroes) + show "P (( *v) A)" + using matrix_vector_mul_linear + by (rule zeroes[where i=i]) + (metis \row i A = 0\ inner_zero_left matrix_vector_mul_component row_def vec_lambda_eta) next fix A :: "real^'n^'n" assume 0: "\i j. i \ j \ A $ i $ j = 0" @@ -246,7 +253,7 @@ = measure lebesgue (cbox a b)" (is "?Q") proof - have lin: "linear ?f" - by (force simp: plus_vec_def scaleR_vec_def algebra_simps intro: linearI) + by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps) show fab: "?f ` cbox a b \ lmeasurable" by (simp add: lin measurable_linear_image_interval) let ?c = "\ i. if i = m then b$m + b$n else b$i" @@ -474,7 +481,7 @@ then have "\ inj f" by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component) have detf: "det (matrix f) = 0" - by (metis "0" \linear f\ invertible_det_nz invertible_right_inverse matrix_right_invertible_surjective matrix_vector_mul surjE vec_component) + using \\ inj f\ det_nz_iff_inj[OF \linear f\] by blast show "f ` S \ lmeasurable \ ?Q f S" proof show "f ` S \ lmeasurable" @@ -500,7 +507,7 @@ assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i" have lin: "linear ?h" - by (simp add: plus_vec_def scaleR_vec_def linearI) + by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def) have meq: "measure lebesgue ((\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i) ` cbox a b) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") @@ -532,7 +539,7 @@ assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. if i = m then v $ m + v $ n else v $ i" have lin: "linear ?h" - by (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff intro: linearI) + by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff) consider "m < n" | " n < m" using \m \ n\ less_linear by blast then have 1: "det(matrix ?h) = 1" @@ -997,11 +1004,11 @@ then show ?thesis proof (rule_tac x="(y - x) /\<^sub>R r" in bexI) have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r" - by (simp add: lin linear_cmul) + by (simp add: lin linear_scale) then have "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) = norm (f' x (y - x) /\<^sub>R r - (f y - f x) /\<^sub>R r)" by (simp add: dist_norm) also have "\ = norm (f' x (y - x) - (f y - f x)) / r" - using \r > 0\ by (simp add: real_vector.scale_right_diff_distrib [symmetric] divide_simps) + using \r > 0\ by (simp add: scale_right_diff_distrib [symmetric] divide_simps) also have "\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x)" using that \r > 0\ False by (simp add: algebra_simps divide_simps dist_norm norm_minus_commute mult_right_mono) also have "\ < k" @@ -1610,8 +1617,9 @@ and lb: "\v. v \ 0 \ (\k>0. \e>0. \x. x \ S - {0} \ norm x < e \ k * norm x \ \v \ x\)" shows "f x = 0" proof - + interpret linear f by fact have "dim {x. f x = 0} \ DIM('a)" - using dim_subset_UNIV by blast + by (rule dim_subset_UNIV) moreover have False if less: "dim {x. f x = 0} < DIM('a)" proof - obtain d where "d \ 0" and d: "\y. f y = 0 \ d \ y = 0" @@ -1662,7 +1670,7 @@ with lim0 \ have "((\x. f x /\<^sub>R norm x) \ (\ \ \)) \ 0" by (force simp: tendsto_at_iff_sequentially) then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ 0" - by (simp add: o_def linear_cmul \linear f\) + by (simp add: o_def scale) qed qed ultimately show False @@ -1671,7 +1679,9 @@ ultimately have dim: "dim {x. f x = 0} = DIM('a)" by force then show ?thesis - by (metis (mono_tags, lifting) UNIV_I assms(1) dim_eq_full linear_eq_0_span mem_Collect_eq) + using dim_eq_full + by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis + mem_Collect_eq module_hom_zero span_base span_raw_def) qed lemma lemma_partial_derivatives: @@ -1950,11 +1960,12 @@ unfolding subspace_def convergent_eq_Cauchy [symmetric] by (force simp: algebra_simps intro: tendsto_intros) then have CA_eq: "?CA = span ?CA" - by (metis span_eq) + by (metis span_eq_iff) also have "\ = UNIV" proof - have "dim ?CA \ CARD('m)" - by (rule dim_subset_UNIV_cart) + using dim_subset_UNIV[of ?CA] + by auto moreover have "False" if less: "dim ?CA < CARD('m)" proof - obtain d where "d \ 0" and d: "\y. y \ span ?CA \ orthogonal d y" @@ -2102,7 +2113,9 @@ have lin_df: "linear (f' x)" and lim_df: "((\y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \ 0) (at x within S)" using \x \ S\ assms by (auto simp: has_derivative_within linear_linear) - moreover have "(matrix (f' x) - B) *v w = 0" for w + moreover + interpret linear "f' x" by fact + have "(matrix (f' x) - B) *v w = 0" for w proof (rule lemma_partial_derivatives [of "( *v) (matrix (f' x) - B)"]) show "linear (( *v) (matrix (f' x) - B))" by (rule matrix_vector_mul_linear) @@ -2124,7 +2137,8 @@ fix k :: "nat" assume "0 \ k" have "0 \ onorm (( *v) (A k - B))" - by (simp add: linear_linear onorm_pos_le matrix_vector_mul_linear) + using matrix_vector_mul_bounded_linear + by (rule onorm_pos_le) then show "norm (onorm (( *v) (A k - B))) \ (\i\UNIV. \j\UNIV. \(A k - B) $ i $ j\)" by (simp add: onorm_le_matrix_component_sum del: vector_minus_component) qed @@ -2174,11 +2188,11 @@ then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" - by (simp add: algebra_simps lin_df linear_diff matrix_vector_mul_linear) + by (simp add: algebra_simps diff lin_df matrix_vector_mul_linear scalar_mult_eq_scaleR) qed qed (use x in \simp; auto simp: not_less\) ultimately have "f' x = ( *v) B" - by (force simp: algebra_simps) + by (force simp: algebra_simps scalar_mult_eq_scaleR) show "matrix (f' x) $ m $ n \ b" proof (rule tendsto_upperbound [of "\i. (A i $ m $ n)" _ sequentially]) show "(\i. A i $ m $ n) \ matrix (f' x) $ m $ n" @@ -2228,11 +2242,11 @@ have "linear (f' x)" using True f has_derivative_linear by blast then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))" - by (metis matrix_vector_mul matrix_vector_mult_diff_rdistrib) + by (simp add: matrix_vector_mult_diff_rdistrib) also have "\ \ (e * norm (y - x)) / 2" proof (rule split246) have "norm ((?A - B) *v (y - x)) / norm (y - x) \ onorm(\x. (?A - B) *v x)" - by (simp add: le_onorm linear_linear matrix_vector_mul_linear) + by (rule le_onorm) auto also have "\ < e/6" by (rule Bo_e6) finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" . @@ -2445,7 +2459,8 @@ moreover have "\t\T -` P. norm (inv T x - t) \ e" proof have "T (inv T x - inv T t) = x - t" - using T linear_diff orthogonal_transformation_def by fastforce + using T linear_diff orthogonal_transformation_def + by (metis (no_types, hide_lams) Tinv) then have "norm (inv T x - inv T t) = norm (x - t)" by (metis T orthogonal_transformation_norm) then show "norm (inv T x - inv T t) \ e" @@ -2473,7 +2488,7 @@ and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain a where "a \ 0" "P \ {x. a \ x = 0}" - using lowdim_subset_hyperplane [of P] P span_inc by auto + using lowdim_subset_hyperplane [of P] P span_base by auto then obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" @@ -2566,7 +2581,8 @@ obtain x where x: "x \ S \ cbox u v" "cbox u v \ ball x (r x)" using \K \ \\ covered uv by blast then have "dim (range (f' x)) < ?n" - using rank_dim_range [of "matrix (f' x)"] lin_f' rank by fastforce + using rank_dim_range [of "matrix (f' x)"] x rank[of x] + by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f') then obtain T where T: "T \ lmeasurable" and subT: "{z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))} \ T" and measT: "?\ T \ (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)" @@ -2586,8 +2602,8 @@ using r [of x y] x \y \ S\ by blast have yx_le: "norm (y - x) \ norm (v - u)" proof (rule norm_le_componentwise_cart) - show "\(y - x) $ i\ \ \(v - u) $ i\" for i - using x y by (force simp: mem_box_cart dest!: spec [where x=i]) + show "norm ((y - x) $ i) \ norm ((v - u) $ i)" for i + using x y by (force simp: mem_box_cart dest!: spec [where x=i]) qed have *: "\norm(y - x - z) \ d; norm z \ B; d \ B\ \ norm(y - x) \ 2 * B" for x y z :: "real^'n::_" and d B @@ -3177,8 +3193,8 @@ if "x \ T" for x proof - have "matrix (h' x) ** matrix (g' (h x)) = mat 1" - using that id matrix_compose - by (metis der_g gh has_derivative_linear left_inverse_linear matrix_id_mat_1) + using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear + by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+ then have "\det (matrix (h' x))\ * \det (matrix (g' (h x)))\ = 1" by (metis abs_1 abs_mult det_I det_mul) then show ?thesis diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Connected.thy --- a/src/HOL/Analysis/Connected.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Connected.thy Wed May 02 13:49:38 2018 +0200 @@ -4141,13 +4141,13 @@ lemma closed_span [iff]: "closed (span s)" for s :: "'a::euclidean_space set" - by (simp add: closed_subspace) + by (simp add: closed_subspace subspace_span) lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") for s :: "'a::euclidean_space set" proof - have "?dc \ ?d" - using closure_minimal[OF span_inc, of s] + using closure_minimal[OF span_superset, of s] using closed_subspace[OF subspace_span, of s] using dim_subset[of "closure s" "span s"] by simp diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Wed May 02 13:49:38 2018 +0200 @@ -28,50 +28,6 @@ done qed -lemma dim_image_eq: - fixes f :: "'n::euclidean_space \ 'm::euclidean_space" - assumes lf: "linear f" - and fi: "inj_on f (span S)" - shows "dim (f ` S) = dim (S::'n::euclidean_space set)" -proof - - obtain B where B: "B \ S" "independent B" "S \ span B" "card B = dim S" - using basis_exists[of S] by auto - then have "span S = span B" - using span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto - then have "independent (f ` B)" - using independent_inj_on_image[of B f] B assms by auto - moreover have "card (f ` B) = card B" - using assms card_image[of f B] subset_inj_on[of f "span S" B] B span_inc by auto - moreover have "(f ` B) \ (f ` S)" - using B by auto - ultimately have "dim (f ` S) \ dim S" - using independent_card_le_dim[of "f ` B" "f ` S"] B by auto - then show ?thesis - using dim_image_le[of f S] assms by auto -qed - -lemma linear_injective_on_subspace_0: - assumes lf: "linear f" - and "subspace S" - shows "inj_on f S \ (\x \ S. f x = 0 \ x = 0)" -proof - - have "inj_on f S \ (\x \ S. \y \ S. f x = f y \ x = y)" - by (simp add: inj_on_def) - also have "\ \ (\x \ S. \y \ S. f x - f y = 0 \ x - y = 0)" - by simp - also have "\ \ (\x \ S. \y \ S. f (x - y) = 0 \ x - y = 0)" - by (simp add: linear_diff[OF lf]) - also have "\ \ (\x \ S. f x = 0 \ x = 0)" - using \subspace S\ subspace_def[of S] subspace_diff[of S] by auto - finally show ?thesis . -qed - -lemma subspace_Inter: "\s \ f. subspace s \ subspace (\f)" - unfolding subspace_def by auto - -lemma span_eq[simp]: "span s = s \ subspace s" - unfolding span_def by (rule hull_eq) (rule subspace_Inter) - lemma substdbasis_expansion_unique: assumes d: "d \ Basis" shows "(\i\d. f i *\<^sub>R i) = (x::'a::euclidean_space) \ @@ -105,20 +61,16 @@ moreover from * have "x = (norm x/e) *\<^sub>R y" by auto ultimately have "x \ span (cball 0 e)" - using span_mul[of y "cball 0 e" "norm x/e"] span_inc[of "cball 0 e"] - by (simp add: span_superset) + using span_scale[of y "cball 0 e" "norm x/e"] + span_superset[of "cball 0 e"] + by (simp add: span_base) } then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" by auto then show ?thesis - using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp add: dim_UNIV) -qed - -lemma indep_card_eq_dim_span: - fixes B :: "'n::euclidean_space set" - assumes "independent B" - shows "finite B \ card B = dim (span B)" - using assms basis_card_eq_dim[of B "span B"] span_inc by auto + using dim_span[of "cball (0 :: 'n::euclidean_space) e"] + by auto +qed lemma sum_not_0: "sum f A \ 0 \ \a \ A. f a \ 0" by (rule ccontr) auto @@ -1218,57 +1170,6 @@ shows "\linear f; inj f\ \ convex (f ` s) \ convex s" by (metis (no_types) convex_linear_image convex_linear_vimage inj_vimage_image_eq) -lemma basis_to_basis_subspace_isomorphism: - assumes s: "subspace (S:: ('n::euclidean_space) set)" - and t: "subspace (T :: ('m::euclidean_space) set)" - and d: "dim S = dim T" - and B: "B \ S" "independent B" "S \ span B" "card B = dim S" - and C: "C \ T" "independent C" "T \ span C" "card C = dim T" - shows "\f. linear f \ f ` B = C \ f ` S = T \ inj_on f S" -proof - - from B independent_bound have fB: "finite B" - by blast - from C independent_bound have fC: "finite C" - by blast - from B(4) C(4) card_le_inj[of B C] d obtain f where - f: "f ` B \ C" "inj_on f B" using \finite B\ \finite C\ by auto - from linear_independent_extend[OF B(2)] obtain g where - g: "linear g" "\x \ B. g x = f x" by blast - from inj_on_iff_eq_card[OF fB, of f] f(2) - have "card (f ` B) = card B" by simp - with B(4) C(4) have ceq: "card (f ` B) = card C" using d - by simp - have "g ` B = f ` B" using g(2) - by (auto simp add: image_iff) - also have "\ = C" using card_subset_eq[OF fC f(1) ceq] . - finally have gBC: "g ` B = C" . - have gi: "inj_on g B" using f(2) g(2) - by (auto simp add: inj_on_def) - note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] - { - fix x y - assume x: "x \ S" and y: "y \ S" and gxy: "g x = g y" - from B(3) x y have x': "x \ span B" and y': "y \ span B" - by blast+ - from gxy have th0: "g (x - y) = 0" - by (simp add: linear_diff[OF g(1)]) - have th1: "x - y \ span B" using x' y' - by (metis span_diff) - have "x = y" using g0[OF th1 th0] by simp - } - then have giS: "inj_on g S" unfolding inj_on_def by blast - from span_subspace[OF B(1,3) s] - have "g ` S = span (g ` B)" - by (simp add: span_linear_image[OF g(1)]) - also have "\ = span C" - unfolding gBC .. - also have "\ = T" - using span_subspace[OF C(1,3) t] . - finally have gS: "g ` S = T" . - from g(1) gS giS gBC show ?thesis - by blast -qed - lemma closure_bounded_linear_image_subset: assumes f: "bounded_linear f" shows "f ` closure S \ closure (f ` S)" @@ -1820,19 +1721,13 @@ assume as: "finite t" "t \ {}" "t \ insert a s" "sum u t = 1" "(\v\t. u v *\<^sub>R v) = x" have "(\x. x - a) ` (t - {a}) \ {x - a |x. x \ s}" using as(3) by auto - then show "\v. x = a + v \ (\S u. finite S \ S \ {x - a |x. x \ s} \ (\v\S. u v *\<^sub>R v) = v)" + then show "\v. x = a + v \ (\S u. v = (\v\S. u v *\<^sub>R v) \ finite S \ S \ {x - a |x. x \ s} )" apply (rule_tac x="x - a" in exI) apply (rule conjI, simp) apply (rule_tac x="(\x. x - a) ` (t - {a})" in exI) apply (rule_tac x="\x. u (x + a)" in exI) - apply (rule conjI) using as(1) apply simp - apply (erule conjI) - using as(1) - apply (simp add: sum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib - sum_subtractf scaleR_left.sum[symmetric] sum_diff1 scaleR_left_diff_distrib) - unfolding as - apply simp - done + by (simp_all add: as sum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib + sum_subtractf scaleR_left.sum[symmetric] sum_diff1 scaleR_left_diff_distrib) qed lemma affine_hull_insert_span: @@ -2113,7 +2008,7 @@ unfolding cone_def by auto lemma subspace_imp_cone: "subspace S \ cone S" - by (simp add: cone_def subspace_mul) + by (simp add: cone_def subspace_scale) subsubsection \Conic hull\ @@ -3153,10 +3048,10 @@ using subspace_imp_affine affine_imp_convex by auto lemma affine_hull_subset_span: "(affine hull s) \ (span s)" - by (metis hull_minimal span_inc subspace_imp_affine subspace_span) + by (metis hull_minimal span_superset subspace_imp_affine subspace_span) lemma convex_hull_subset_span: "(convex hull s) \ (span s)" - by (metis hull_minimal span_inc subspace_imp_convex subspace_span) + by (metis hull_minimal span_superset subspace_imp_convex subspace_span) lemma convex_hull_subset_affine_hull: "(convex hull s) \ (affine hull s)" by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset) @@ -3289,7 +3184,7 @@ apply (rule subset_le_dim) unfolding subset_eq using \a\s\ - apply (auto simp add:span_superset span_diff) + apply (auto simp add:span_base span_diff) done also have "\ < dim s + 1" by auto also have "\ \ card (s - {a})" @@ -3660,45 +3555,7 @@ then show ?thesis by auto qed -lemma independent_finite: - fixes B :: "'n::euclidean_space set" - assumes "independent B" - shows "finite B" - using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms - by auto - -lemma subspace_dim_equal: - assumes "subspace (S :: ('n::euclidean_space) set)" - and "subspace T" - and "S \ T" - and "dim S \ dim T" - shows "S = T" -proof - - obtain B where B: "B \ S" "independent B \ S \ span B" "card B = dim S" - using basis_exists[of S] by auto - then have "span B \ S" - using span_mono[of B S] span_eq[of S] assms by metis - then have "span B = S" - using B by auto - have "dim S = dim T" - using assms dim_subset[of S T] by auto - then have "T \ span B" - using card_eq_dim[of B T] B independent_finite assms by auto - then show ?thesis - using assms \span B = S\ by auto -qed - -corollary dim_eq_span: - fixes S :: "'a::euclidean_space set" - shows "\S \ T; dim T \ dim S\ \ span S = span T" -by (simp add: span_mono subspace_dim_equal subspace_span) - -lemma dim_eq_full: - fixes S :: "'a :: euclidean_space set" - shows "dim S = DIM('a) \ span S = UNIV" -apply (rule iffI) - apply (metis dim_eq_span dim_subset_UNIV span_Basis span_span subset_UNIV) -by (metis dim_UNIV dim_span) +lemmas independent_finite = independent_imp_finite lemma span_substd_basis: assumes d: "d \ Basis" @@ -3710,9 +3567,10 @@ moreover have s: "subspace ?B" using subspace_substandard[of "\i. i \ d"] . ultimately have "span d \ ?B" - using span_mono[of d "?B"] span_eq[of "?B"] by blast + using span_mono[of d "?B"] span_eq_iff[of "?B"] by blast moreover have *: "card d \ dim (span d)" - using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] span_inc[of d] + using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] + span_superset[of d] by auto moreover from * have "dim ?B \ dim (span d)" using dim_substandard[OF assms] by auto @@ -3727,7 +3585,7 @@ f ` span B = {x. \i\Basis. i \ d \ x \ i = 0} \ inj_on f (span B) \ d \ Basis" proof - have B: "card B = dim B" - using dim_unique[of B B "card B"] assms span_inc[of B] by auto + using dim_unique[of B B "card B"] assms span_superset[of B] by auto have "dim B \ card (Basis :: 'a set)" using dim_subset_UNIV[of B] by simp from ex_card[OF this] obtain d :: "'a set" where d: "d \ Basis" and t: "card d = dim B" @@ -3738,13 +3596,13 @@ apply (rule subspace_span) apply (rule subspace_substandard) defer - apply (rule span_inc) + apply (rule span_superset) apply (rule assms) defer unfolding dim_span[of B] apply(rule B) unfolding span_substd_basis[OF d, symmetric] - apply (rule span_inc) + apply (rule span_superset) apply (rule independent_substdbasis[OF d]) apply rule apply assumption @@ -4032,7 +3890,7 @@ have "a \ w = a \ c" if "span ((+) (- c) ` S) \ {x. a \ x = 0}" "w \ S" for w proof - have "w-c \ span ((+) (- c) ` S)" - by (simp add: span_superset \w \ S\) + by (simp add: span_base \w \ S\) with that have "w-c \ {x. a \ x = 0}" by blast then show ?thesis @@ -4052,10 +3910,10 @@ have 1: "{b - a|b. b \ (S - {a})} \ {x - a|x. x \ S}" by auto have 2: "x - a \ span {b - a |b. b \ S - {a}}" if "x \ S" for x proof (cases "x = a") - case True then show ?thesis by simp + case True then show ?thesis by (simp add: span_clauses) next case False then show ?thesis - using assms by (blast intro: span_superset that) + using assms by (blast intro: span_base that) qed have "\ affine_dependent (insert a S)" by (simp add: assms insert_absorb) @@ -5072,7 +4930,7 @@ then have "y \ S" using assms y hull_subset[of S] affine_hull_subset_span inj_on_image_mem_iff [OF \inj_on f (span S)\] - by (metis Int_iff span_inc subsetCE) + by (metis Int_iff span_superset subsetCE) } then have "z \ f ` (rel_interior S)" using mem_rel_interior_cball[of x S] \e > 0\ x by auto @@ -5107,7 +4965,7 @@ using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto moreover have *: "x - xy \ span S" using subspace_diff[of "span S" x xy] subspace_span \x \ S\ xy - affine_hull_subset_span[of S] span_inc + affine_hull_subset_span[of S] span_superset by auto moreover from * have "e1 * norm (x - xy) \ norm (f (x - xy))" using e1 by auto @@ -6070,7 +5928,7 @@ proof - from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] obtain a where "a \ 0" "\x\{x - y |x y. x \ t \ y \ s}. 0 \ inner a x" - using assms(3-5) by fastforce + using assms(3-5) by force then have *: "\x y. x \ t \ y \ s \ inner a y \ inner a x" by (force simp add: inner_diff) then have bdd: "bdd_above (((\) a)`s)" @@ -6152,7 +6010,7 @@ lemma convex_hull_scaling: "convex hull ((\x. c *\<^sub>R x) ` s) = (\x. c *\<^sub>R x) ` (convex hull s)" - using linear_scaleR by (rule convex_hull_linear_image [symmetric]) + by (simp add: convex_hull_linear_image) lemma convex_hull_affinity: "convex hull ((\x. a + c *\<^sub>R x) ` s) = (\x. a + c *\<^sub>R x) ` (convex hull s)" @@ -6730,8 +6588,10 @@ shows "connected s \ convex s" proof - obtain f:: "'a \ real" where linf: "linear f" and "inj f" - using subspace_isomorphism [where 'a = 'a and 'b = real] - by (metis DIM_real dim_UNIV subspace_UNIV assms) + using subspace_isomorphism[OF subspace_UNIV subspace_UNIV, + where 'a='a and 'b=real] + unfolding Euclidean_Space.dim_UNIV + by (auto simp: assms) then have "f -` (f ` s) = s" by (simp add: inj_vimage_image_eq) then show ?thesis @@ -6739,7 +6599,7 @@ qed lemma is_interval_cball_1[intro, simp]: "is_interval (cball a b)" for a b::real - by (auto simp: is_interval_convex_1 convex_cball) + by (simp add: is_interval_convex_1) subsection%unimportant \Another intermediate value theorem formulation\ @@ -6909,7 +6769,7 @@ also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` (convex hull {0, 1}))" by (simp only: convex_hull_eq_real_cbox zero_le_one) also have "\ = (\i\Basis. convex hull ((\x. x *\<^sub>R i) ` {0, 1}))" - by (simp only: convex_hull_linear_image linear_scaleR_left) + by (simp add: convex_hull_linear_image) also have "\ = convex hull (\i\Basis. (\x. x *\<^sub>R i) ` {0, 1})" by (simp only: convex_hull_set_sum) also have "\ = convex hull {x. \i\Basis. x\i \ {0, 1}}" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Derivative.thy Wed May 02 13:49:38 2018 +0200 @@ -400,7 +400,7 @@ apply (rule linear_eq_stdbasis) unfolding linear_conv_bounded_linear apply (rule as(1,2)[THEN conjunct1])+ - proof (rule, rule ccontr) + proof (rule ccontr) fix i :: 'a assume i: "i \ Basis" define e where "e = norm (f' i - f'' i)" @@ -1605,28 +1605,6 @@ We could put \f' \ g = I\ but this happens to fit with the minimal linear algebra theory I've set up so far.\ -(* move before left_inverse_linear in Euclidean_Space*) - -lemma right_inverse_linear: - fixes f :: "'a::euclidean_space \ 'a" - assumes lf: "linear f" - and gf: "f \ g = id" - shows "linear g" -proof - - from gf have fi: "surj f" - by (auto simp add: surj_def o_def id_def) metis - from linear_surjective_isomorphism[OF lf fi] - obtain h:: "'a \ 'a" where h: "linear h" "\x. h (f x) = x" "\x. f (h x) = x" - by blast - have "h = g" - apply (rule ext) - using gf h(2,3) - apply (simp add: o_def id_def fun_eq_iff) - apply metis - done - with h(1) show ?thesis by blast -qed - lemma has_derivative_inverse_strong: fixes f :: "'n::euclidean_space \ 'n" assumes "open s" @@ -2778,7 +2756,7 @@ "\\<^sub>F n in sequentially. ((\x. \i A) ?F" by eventually_elim (auto intro!: tendsto_eq_intros - simp: power_0_left if_distrib cond_application_beta sum.delta + simp: power_0_left if_distrib if_distribR sum.delta cong: if_cong) ultimately have [tendsto_intros]: "((\x. \i. ?e i x) \ A) ?F" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Determinants.thy --- a/src/HOL/Analysis/Determinants.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Determinants.thy Wed May 02 13:49:38 2018 +0200 @@ -246,52 +246,106 @@ by simp qed -lemma det_identical_rows: - fixes A :: "'a::linordered_idom^'n^'n" - assumes ij: "i \ j" - and r: "row i A = row j A" +lemma det_identical_columns: + fixes A :: "'a::comm_ring_1^'n^'n" + assumes jk: "j \ k" + and r: "column j A = column k A" shows "det A = 0" -proof- - have tha: "\(a::'a) b. a = b \ b = - a \ a = 0" - by simp - have th1: "of_int (-1) = - 1" by simp - let ?p = "Fun.swap i j id" - let ?A = "\ i. A $ ?p i" - from r have "A = ?A" by (simp add: vec_eq_iff row_def Fun.swap_def) - then have "det A = det ?A" by simp - moreover have "det A = - det ?A" - by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1) - ultimately show "det A = 0" by (metis tha) +proof - + let ?U="UNIV::'n set" + let ?t_jk="Fun.swap j k id" + let ?PU="{p. p permutes ?U}" + let ?S1="{p. p\?PU \ evenperm p}" + let ?S2="{(?t_jk \ p) |p. p \?S1}" + let ?f="\p. of_int (sign p) * (\i\UNIV. A $ i $ p i)" + let ?g="\p. ?t_jk \ p" + have g_S1: "?S2 = ?g` ?S1" by auto + have inj_g: "inj_on ?g ?S1" + proof (unfold inj_on_def, auto) + fix x y assume x: "x permutes ?U" and even_x: "evenperm x" + and y: "y permutes ?U" and even_y: "evenperm y" and eq: "?t_jk \ x = ?t_jk \ y" + show "x = y" by (metis (hide_lams, no_types) comp_assoc eq id_comp swap_id_idempotent) + qed + have tjk_permutes: "?t_jk permutes ?U" unfolding permutes_def swap_id_eq by (auto,metis) + have tjk_eq: "\i l. A $ i $ ?t_jk l = A $ i $ l" + using r jk + unfolding column_def vec_eq_iff swap_id_eq by fastforce + have sign_tjk: "sign ?t_jk = -1" using sign_swap_id[of j k] jk by auto + {fix x + assume x: "x\ ?S1" + have "sign (?t_jk \ x) = sign (?t_jk) * sign x" + by (metis (lifting) finite_class.finite_UNIV mem_Collect_eq + permutation_permutes permutation_swap_id sign_compose x) + also have "... = - sign x" using sign_tjk by simp + also have "... \ sign x" unfolding sign_def by simp + finally have "sign (?t_jk \ x) \ sign x" and "(?t_jk \ x) \ ?S2" + by (auto, metis (lifting, full_types) mem_Collect_eq x) + } + hence disjoint: "?S1 \ ?S2 = {}" by (auto, metis sign_def) + have PU_decomposition: "?PU = ?S1 \ ?S2" + proof (auto) + fix x + assume x: "x permutes ?U" and "\p. p permutes ?U \ x = Fun.swap j k id \ p \ \ evenperm p" + from this obtain p where p: "p permutes UNIV" and x_eq: "x = Fun.swap j k id \ p" + and odd_p: "\ evenperm p" + by (metis (no_types) comp_assoc id_comp inv_swap_id permutes_compose + permutes_inv_o(1) tjk_permutes) + thus "evenperm x" + by (metis evenperm_comp evenperm_swap finite_class.finite_UNIV + jk permutation_permutes permutation_swap_id) + next + fix p assume p: "p permutes ?U" + show "Fun.swap j k id \ p permutes UNIV" by (metis p permutes_compose tjk_permutes) + qed + have "sum ?f ?S2 = sum ((\p. of_int (sign p) * (\i\UNIV. A $ i $ p i)) + \ (\) (Fun.swap j k id)) {p \ {p. p permutes UNIV}. evenperm p}" + unfolding g_S1 by (rule sum.reindex[OF inj_g]) + also have "... = sum (\p. of_int (sign (?t_jk \ p)) * (\i\UNIV. A $ i $ p i)) ?S1" + unfolding o_def by (rule sum.cong, auto simp add: tjk_eq) + also have "... = sum (\p. - ?f p) ?S1" + proof (rule sum.cong, auto) + fix x assume x: "x permutes ?U" + and even_x: "evenperm x" + hence perm_x: "permutation x" and perm_tjk: "permutation ?t_jk" + using permutation_permutes[of x] permutation_permutes[of ?t_jk] permutation_swap_id + by (metis finite_code)+ + have "(sign (?t_jk \ x)) = - (sign x)" + unfolding sign_compose[OF perm_tjk perm_x] sign_tjk by auto + thus "of_int (sign (?t_jk \ x)) * (\i\UNIV. A $ i $ x i) + = - (of_int (sign x) * (\i\UNIV. A $ i $ x i))" + by auto + qed + also have "...= - sum ?f ?S1" unfolding sum_negf .. + finally have *: "sum ?f ?S2 = - sum ?f ?S1" . + have "det A = (\p | p permutes UNIV. of_int (sign p) * (\i\UNIV. A $ i $ p i))" + unfolding det_def .. + also have "...= sum ?f ?S1 + sum ?f ?S2" + by (subst PU_decomposition, rule sum.union_disjoint[OF _ _ disjoint], auto) + also have "...= sum ?f ?S1 - sum ?f ?S1 " unfolding * by auto + also have "...= 0" by simp + finally show "det A = 0" by simp qed -lemma det_identical_columns: - fixes A :: "'a::linordered_idom^'n^'n" +lemma det_identical_rows: + fixes A :: "'a::comm_ring_1^'n^'n" assumes ij: "i \ j" - and r: "column i A = column j A" + and r: "row i A = row j A" shows "det A = 0" apply (subst det_transpose[symmetric]) - apply (rule det_identical_rows[OF ij]) - apply (metis row_transpose r) + apply (rule det_identical_columns[OF ij]) + apply (metis column_transpose r) done lemma det_zero_row: - fixes A :: "'a::{idom, ring_char_0}^'n^'n" - assumes r: "row i A = 0" - shows "det A = 0" - using r - apply (simp add: row_def det_def vec_eq_iff) - apply (rule sum.neutral) - apply (auto simp: sign_nz) - done + fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m" + shows "row i A = 0 \ det A = 0" and "row j F = 0 \ det F = 0" + by (force simp add: row_def det_def vec_eq_iff sign_nz intro!: sum.neutral)+ lemma det_zero_column: - fixes A :: "'a::{idom,ring_char_0}^'n^'n" - assumes r: "column i A = 0" - shows "det A = 0" - apply (subst det_transpose[symmetric]) - apply (rule det_zero_row [of i]) - apply (metis row_transpose r) - done + fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m" + shows "column i A = 0 \ det A = 0" and "column j F = 0 \ det F = 0" + unfolding atomize_conj atomize_imp + by (metis det_transpose det_zero_row row_transpose) lemma det_row_add: fixes a b c :: "'n::finite \ _ ^ 'n" @@ -404,7 +458,7 @@ done lemma det_row_operation: - fixes A :: "'a::linordered_idom^'n^'n" + fixes A :: "'a::{comm_ring_1}^'n^'n" assumes ij: "i \ j" shows "det (\ k. if k = i then row i A + c *s row j A else row k A) = det A" proof - @@ -418,50 +472,35 @@ qed lemma det_row_span: - fixes A :: "real^'n^'n" - assumes x: "x \ span {row j A |j. j \ i}" + fixes A :: "'a::{field}^'n^'n" + assumes x: "x \ vec.span {row j A |j. j \ i}" shows "det (\ k. if k = i then row i A + x else row k A) = det A" -proof - - let ?U = "UNIV :: 'n set" - let ?S = "{row j A |j. j \ i}" + using x +proof (induction rule: vec.span_induct_alt) + case 1 + then show ?case + by (rule cong[of det, OF refl]) (vector row_def) +next + case (2 c z y) + note zS = 2(1) + and Py = 2(2) let ?d = "\x. det (\ k. if k = i then x else row k A)" let ?P = "\x. ?d (row i A + x) = det A" - { - fix k - have "(if k = i then row i A + 0 else row k A) = row k A" - by simp - } - then have P0: "?P 0" - apply - - apply (rule cong[of det, OF refl]) + from zS obtain j where j: "z = row j A" "i \ j" + by blast + let ?w = "row i A + y" + have th0: "row i A + (c*s z + y) = ?w + c*s z" + by vector + have thz: "?d z = 0" + apply (rule det_identical_rows[OF j(2)]) + using j apply (vector row_def) done - moreover - { - fix c z y - assume zS: "z \ ?S" and Py: "?P y" - from zS obtain j where j: "z = row j A" "i \ j" - by blast - let ?w = "row i A + y" - have th0: "row i A + (c*s z + y) = ?w + c*s z" - by vector - have thz: "?d z = 0" - apply (rule det_identical_rows[OF j(2)]) - using j - apply (vector row_def) - done - have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" - unfolding th0 .. - then have "?P (c*s z + y)" - unfolding thz Py det_row_mul[of i] det_row_add[of i] - by simp - } - ultimately show ?thesis - apply - - apply (rule span_induct_alt[of ?P ?S, OF P0, folded scalar_mult_eq_scaleR]) - apply blast - apply (rule x) - done + have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" + unfolding th0 .. + then show ?case + unfolding thz Py det_row_mul[of i] det_row_add[of i] + by simp qed lemma matrix_id [simp]: "det (matrix id) = 1" @@ -469,7 +508,7 @@ lemma det_matrix_scaleR [simp]: "det (matrix ((( *\<^sub>R) r)) :: real^'n^'n) = r ^ CARD('n::finite)" apply (subst det_diagonal) - apply (auto simp: matrix_def mat_def prod_constant) + apply (auto simp: matrix_def mat_def) apply (simp add: cart_eq_inner_axis inner_axis_axis) done @@ -479,13 +518,13 @@ \ lemma det_dependent_rows: - fixes A:: "real^'n^'n" - assumes d: "dependent (rows A)" + fixes A:: "'a::{field}^'n^'n" + assumes d: "vec.dependent (rows A)" shows "det A = 0" proof - let ?U = "UNIV :: 'n set" - from d obtain i where i: "row i A \ span (rows A - {row i A})" - unfolding dependent_def rows_def by blast + from d obtain i where i: "row i A \ vec.span (rows A - {row i A})" + unfolding vec.dependent_def rows_def by blast { fix j k assume jk: "j \ k" and c: "row j A = row k A" @@ -494,25 +533,25 @@ moreover { assume H: "\ i j. i \ j \ row i A \ row j A" - have th0: "- row i A \ span {row j A|j. j \ i}" - apply (rule span_neg) + have th0: "- row i A \ vec.span {row j A|j. j \ i}" + apply (rule vec.span_neg) apply (rule set_rev_mp) apply (rule i) - apply (rule span_mono) + apply (rule vec.span_mono) using H i apply (auto simp add: rows_def) done from det_row_span[OF th0] have "det A = det (\ k. if k = i then 0 *s 1 else row k A)" unfolding right_minus vector_smult_lzero .. - with det_row_mul[of i "0::real" "\i. 1"] + with det_row_mul[of i "0::'a" "\i. 1"] have "det A = 0" by simp } ultimately show ?thesis by blast qed lemma det_dependent_columns: - assumes d: "dependent (columns (A::real^'n^'n))" + assumes d: "vec.dependent (columns (A::real^'n^'n))" shows "det A = 0" by (metis d det_dependent_rows rows_transpose det_transpose) @@ -647,7 +686,7 @@ qed rule lemma det_mul: - fixes A B :: "'a::linordered_idom^'n^'n" + fixes A B :: "'a::comm_ring_1^'n^'n" shows "det (A ** B) = det A * det B" proof - let ?U = "UNIV :: 'n set" @@ -767,30 +806,21 @@ finally show ?thesis unfolding th2 . qed + subsection \Relation to invertibility.\ -lemma invertible_left_inverse: - fixes A :: "real^'n^'n" - shows "invertible A \ (\(B::real^'n^'n). B** A = mat 1)" - by (metis invertible_def matrix_left_right_inverse) - -lemma invertible_right_inverse: - fixes A :: "real^'n^'n" - shows "invertible A \ (\(B::real^'n^'n). A** B = mat 1)" - by (metis invertible_def matrix_left_right_inverse) - lemma invertible_det_nz: - fixes A::"real ^'n^'n" + fixes A::"'a::{field}^'n^'n" shows "invertible A \ det A \ 0" proof - { assume "invertible A" - then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1" + then obtain B :: "'a^'n^'n" where B: "A ** B = mat 1" unfolding invertible_right_inverse by blast - then have "det (A ** B) = det (mat 1 :: real ^'n^'n)" + then have "det (A ** B) = det (mat 1 :: 'a^'n^'n)" by simp then have "det A \ 0" - by (simp add: det_mul det_I) algebra + by (simp add: det_mul) algebra } moreover { @@ -804,37 +834,41 @@ unfolding invertible_right_inverse unfolding matrix_right_invertible_independent_rows by blast - have *: "\(a::real^'n) b. a + b = 0 \ -a = b" + have *: "\(a::'a^'n) b. a + b = 0 \ -a = b" apply (drule_tac f="(+) (- a)" in cong[OF refl]) apply (simp only: ab_left_minus add.assoc[symmetric]) apply simp done + from c ci have thr0: "- row i A = sum (\j. (1/ c i) *s (c j *s row j A)) (?U - {i})" + unfolding sum.remove[OF fU iU] sum_cmul + apply - apply (rule vector_mul_lcancel_imp[OF ci]) - using c ci unfolding sum.remove[OF fU iU] sum_cmul - apply (auto simp add: field_simps *) + apply (auto simp add: field_simps) + unfolding * + apply rule done - have thr: "- row i A \ span {row j A| j. j \ i}" + have thr: "- row i A \ vec.span {row j A| j. j \ i}" unfolding thr0 - apply (rule span_sum) + apply (rule vec.span_sum) apply simp - apply (rule span_mul [where 'a="real^'n"]) - apply (rule span_superset) + apply (rule vec.span_scale[folded scalar_mult_eq_scaleR])+ + apply (rule vec.span_base) apply auto done - let ?B = "(\ k. if k = i then 0 else row k A) :: real ^'n^'n" + let ?B = "(\ k. if k = i then 0 else row k A) :: 'a^'n^'n" have thrb: "row i ?B = 0" using iU by (vector row_def) have "det A = 0" unfolding det_row_span[OF thr, symmetric] right_minus - unfolding det_zero_row[OF thrb] .. + unfolding det_zero_row(2)[OF thrb] .. } ultimately show ?thesis by blast qed -lemma det_nz_iff_inj: - fixes f :: "real^'n \ real^'n" - assumes "linear f" +lemma det_nz_iff_inj_gen: + fixes f :: "'a::field^'n \ 'a::field^'n" + assumes "Vector_Spaces.linear ( *s) ( *s) f" shows "det (matrix f) \ 0 \ inj f" proof assume "det (matrix f) \ 0" @@ -843,10 +877,17 @@ next assume "inj f" show "det (matrix f) \ 0" - using linear_injective_left_inverse [OF assms \inj f\] - by (metis assms invertible_det_nz invertible_left_inverse matrix_compose matrix_id_mat_1) + using vec.linear_injective_left_inverse [OF assms \inj f\] + by (metis assms invertible_det_nz invertible_left_inverse matrix_compose_gen matrix_id_mat_1) qed +lemma det_nz_iff_inj: + fixes f :: "real^'n \ real^'n" + assumes "linear f" + shows "det (matrix f) \ 0 \ inj f" + using det_nz_iff_inj_gen[of f] assms + unfolding linear_matrix_vector_mul_eq . + lemma det_eq_0_rank: fixes A :: "real^'n^'n" shows "det A = 0 \ rank A < CARD('n)" @@ -855,71 +896,88 @@ subsubsection\Invertibility of matrices and corresponding linear functions\ -lemma matrix_left_invertible: - fixes f :: "real^'m \ real^'n" - assumes "linear f" - shows "((\B. B ** matrix f = mat 1) \ (\g. linear g \ g \ f = id))" +lemma matrix_left_invertible_gen: + fixes f :: "'a::field^'m \ 'a::field^'n" + assumes "Vector_Spaces.linear ( *s) ( *s) f" + shows "((\B. B ** matrix f = mat 1) \ (\g. Vector_Spaces.linear ( *s) ( *s) g \ g \ f = id))" proof safe fix B assume 1: "B ** matrix f = mat 1" - show "\g. linear g \ g \ f = id" + show "\g. Vector_Spaces.linear ( *s) ( *s) g \ g \ f = id" proof (intro exI conjI) - show "linear (\y. B *v y)" - by (simp add: matrix_vector_mul_linear) + show "Vector_Spaces.linear ( *s) ( *s) (\y. B *v y)" + by (simp add:) show "(( *v) B) \ f = id" unfolding o_def - by (metis assms 1 eq_id_iff matrix_vector_mul matrix_vector_mul_assoc matrix_vector_mul_lid) + by (metis assms 1 eq_id_iff matrix_vector_mul(1) matrix_vector_mul_assoc matrix_vector_mul_lid) qed next fix g - assume "linear g" "g \ f = id" + assume "Vector_Spaces.linear ( *s) ( *s) g" "g \ f = id" then have "matrix g ** matrix f = mat 1" - by (metis assms matrix_compose matrix_id_mat_1) + by (metis assms matrix_compose_gen matrix_id_mat_1) then show "\B. B ** matrix f = mat 1" .. qed -lemma matrix_right_invertible: - fixes f :: "real^'m \ real^'n" - assumes "linear f" - shows "((\B. matrix f ** B = mat 1) \ (\g. linear g \ f \ g = id))" +lemma matrix_left_invertible: + "linear f \ ((\B. B ** matrix f = mat 1) \ (\g. linear g \ g \ f = id))" for f::"real^'m \ real^'n" + using matrix_left_invertible_gen[of f] + by (auto simp: linear_matrix_vector_mul_eq) + +lemma matrix_right_invertible_gen: + fixes f :: "'a::field^'m \ 'a^'n" + assumes "Vector_Spaces.linear ( *s) ( *s) f" + shows "((\B. matrix f ** B = mat 1) \ (\g. Vector_Spaces.linear ( *s) ( *s) g \ f \ g = id))" proof safe fix B assume 1: "matrix f ** B = mat 1" - show "\g. linear g \ f \ g = id" + show "\g. Vector_Spaces.linear ( *s) ( *s) g \ f \ g = id" proof (intro exI conjI) - show "linear (( *v) B)" - by (simp add: matrix_vector_mul_linear) + show "Vector_Spaces.linear ( *s) ( *s) (( *v) B)" + by (simp add: ) show "f \ ( *v) B = id" - by (metis 1 assms comp_apply eq_id_iff linear_id matrix_id_mat_1 matrix_vector_mul_assoc matrix_works) + using 1 assms comp_apply eq_id_iff vec.linear_id matrix_id_mat_1 matrix_vector_mul_assoc matrix_works + by (metis (no_types, hide_lams)) qed next fix g - assume "linear g" and "f \ g = id" + assume "Vector_Spaces.linear ( *s) ( *s) g" and "f \ g = id" then have "matrix f ** matrix g = mat 1" - by (metis assms matrix_compose matrix_id_mat_1) + by (metis assms matrix_compose_gen matrix_id_mat_1) then show "\B. matrix f ** B = mat 1" .. qed +lemma matrix_right_invertible: + "linear f \ ((\B. matrix f ** B = mat 1) \ (\g. linear g \ f \ g = id))" for f::"real^'m \ real^'n" + using matrix_right_invertible_gen[of f] + by (auto simp: linear_matrix_vector_mul_eq) + +lemma matrix_invertible_gen: + fixes f :: "'a::field^'m \ 'a::field^'n" + assumes "Vector_Spaces.linear ( *s) ( *s) f" + shows "invertible (matrix f) \ (\g. Vector_Spaces.linear ( *s) ( *s) g \ f \ g = id \ g \ f = id)" + (is "?lhs = ?rhs") +proof + assume ?lhs then show ?rhs + by (metis assms invertible_def left_right_inverse_eq matrix_left_invertible_gen matrix_right_invertible_gen) +next + assume ?rhs then show ?lhs + by (metis assms invertible_def matrix_compose_gen matrix_id_mat_1) +qed + lemma matrix_invertible: - fixes f :: "real^'m \ real^'n" - assumes "linear f" - shows "invertible (matrix f) \ (\g. linear g \ f \ g = id \ g \ f = id)" - (is "?lhs = ?rhs") -proof - assume ?lhs then show ?rhs - by (metis assms invertible_def left_right_inverse_eq matrix_left_invertible matrix_right_invertible) -next - assume ?rhs then show ?lhs - by (metis assms invertible_def matrix_compose matrix_id_mat_1) -qed + "linear f \ invertible (matrix f) \ (\g. linear g \ f \ g = id \ g \ f = id)" + for f::"real^'m \ real^'n" + using matrix_invertible_gen[of f] + by (auto simp: linear_matrix_vector_mul_eq) lemma invertible_eq_bij: - fixes m :: "real^'m^'n" + fixes m :: "'a::field^'m^'n" shows "invertible m \ bij (( *v) m)" - using matrix_invertible [OF matrix_vector_mul_linear] o_bij - apply (auto simp: bij_betw_def) - by (metis left_right_inverse_eq linear_injective_left_inverse [OF matrix_vector_mul_linear] - linear_surjective_right_inverse[OF matrix_vector_mul_linear]) + using matrix_invertible_gen[OF matrix_vector_mul_linear_gen, of m, simplified matrix_of_matrix_vector_mul] + by (metis bij_betw_def left_right_inverse_eq matrix_vector_mul_linear_gen o_bij + vec.linear_injective_left_inverse vec.linear_surjective_right_inverse) + subsection \Cramer's rule.\ @@ -947,9 +1005,9 @@ by simp have thd0: "det (\ i. if i = k then row k A + (\i \ ?Uk. x $ i *s row i A) else row i A) = det A" apply (rule det_row_span) - apply (rule span_sum) - apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR]) - apply (rule span_superset) + apply (rule vec.span_sum) + apply (rule vec.span_scale) + apply (rule vec.span_base) apply auto done show "?lhs = x$k * det A" @@ -972,7 +1030,7 @@ proof - let ?U = "UNIV :: 'n set" have *: "\c. sum (\i. c i *s row i (transpose A)) ?U = sum (\i. c i *s column i A) ?U" - by (auto simp add: row_transpose intro: sum.cong) + by (auto intro: sum.cong) show ?thesis unfolding matrix_mult_sum unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric] @@ -992,7 +1050,7 @@ unfolding invertible_det_nz[symmetric] invertible_def by blast have "(A ** B) *v b = b" - by (simp add: B matrix_vector_mul_lid) + by (simp add: B) then have "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc) then have xe: "\x. A *v x = b" @@ -1021,7 +1079,7 @@ apply auto apply (erule_tac x=v in allE)+ apply (simp add: norm_eq_sqrt_inner) - apply (simp add: dot_norm linear_add[symmetric]) + apply (simp add: dot_norm linear_add[symmetric]) done lemma orthogonal_transformation_id [simp]: "orthogonal_transformation (\x. x)" @@ -1033,7 +1091,7 @@ lemma orthogonal_transformation_compose: "\orthogonal_transformation f; orthogonal_transformation g\ \ orthogonal_transformation(f \ g)" - by (simp add: orthogonal_transformation_def linear_compose) + by (auto simp add: orthogonal_transformation_def linear_compose) lemma orthogonal_transformation_neg: "orthogonal_transformation(\x. -(f x)) \ orthogonal_transformation f" @@ -1074,7 +1132,7 @@ by (metis matrix_left_right_inverse orthogonal_matrix_def) lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)" - by (simp add: orthogonal_matrix_def transpose_mat matrix_mul_lid) + by (simp add: orthogonal_matrix_def) lemma orthogonal_matrix_mul: fixes A :: "real ^'n^'n" @@ -1085,7 +1143,7 @@ unfolding orthogonal_matrix matrix_transpose_mul apply (subst matrix_mul_assoc) apply (subst matrix_mul_assoc[symmetric]) - apply (simp add: matrix_mul_rid) + apply (simp add: ) done lemma orthogonal_transformation_matrix: @@ -1100,8 +1158,9 @@ let ?m1 = "mat 1 :: real ^'n^'n" { assume ot: ?ot - from ot have lf: "linear f" and fd: "\v w. f v \ f w = v \ w" - unfolding orthogonal_transformation_def orthogonal_matrix by blast+ + from ot have lf: "Vector_Spaces.linear ( *s) ( *s) f" and fd: "\v w. f v \ f w = v \ w" + unfolding orthogonal_transformation_def orthogonal_matrix linear_def scalar_mult_eq_scaleR + by blast+ { fix i j let ?A = "transpose ?mf ** ?mf" @@ -1118,20 +1177,21 @@ unfolding orthogonal_matrix by vector with lf have ?rhs + unfolding linear_def scalar_mult_eq_scaleR by blast } moreover { - assume lf: "linear f" and om: "orthogonal_matrix ?mf" + assume lf: "Vector_Spaces.linear ( *s) ( *s) f" and om: "orthogonal_matrix ?mf" from lf om have ?lhs apply (simp only: orthogonal_matrix_def norm_eq orthogonal_transformation) apply (simp only: matrix_works[OF lf, symmetric]) apply (subst dot_matrix_vector_mul) - apply (simp add: dot_matrix_product matrix_mul_lid) + apply (simp add: dot_matrix_product linear_def scalar_mult_eq_scaleR) done } ultimately show ?thesis - by blast + by (auto simp: linear_def scalar_mult_eq_scaleR) qed lemma det_orthogonal_matrix: @@ -1159,7 +1219,7 @@ then have "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)" by simp then have "det Q * det Q = 1" - by (simp add: det_mul det_I det_transpose) + by (simp add: det_mul) then show ?thesis unfolding th . qed @@ -1266,9 +1326,9 @@ proof - obtain S where "a \ S" "pairwise orthogonal S" and noS: "\x. x \ S \ norm x = 1" and "independent S" "card S = CARD('n)" "span S = UNIV" - using vector_in_orthonormal_basis assms by force - with independent_imp_finite obtain f0 where "bij_betw f0 (UNIV::'n set) S" - by (metis finite_class.finite_UNIV finite_same_card_bij) + using vector_in_orthonormal_basis assms by (force simp: ) + then obtain f0 where "bij_betw f0 (UNIV::'n set) S" + by (metis finite_class.finite_UNIV finite_same_card_bij finiteI_independent) then obtain f where f: "bij_betw f (UNIV::'n set) S" and a: "a = f k" using bij_swap_iff [of k "inv f0 a" f0] by (metis UNIV_I \a \ S\ bij_betw_inv_into_right bij_betw_swap_iff swap_apply1) @@ -1299,7 +1359,8 @@ show thesis proof show "orthogonal_transformation ?f" - by (simp add: AB orthogonal_matrix_mul matrix_vector_mul_linear orthogonal_transformation_matrix) + by (subst orthogonal_transformation_matrix) + (auto simp: AB orthogonal_matrix_mul) next show "?f a = b" using \orthogonal_matrix A\ unfolding orthogonal_matrix_def @@ -1321,10 +1382,10 @@ by (auto intro: orthogonal_transformation_exists_1 [of "a /\<^sub>R norm a" "b /\<^sub>R norm b"]) show ?thesis proof - have "linear f" + interpret linear f using f by (simp add: orthogonal_transformation_linear) - then have "f a /\<^sub>R norm a = f (a /\<^sub>R norm a)" - by (simp add: linear_cmul [of f]) + have "f a /\<^sub>R norm a = f (a /\<^sub>R norm a)" + by (simp add: scale) also have "\ = b /\<^sub>R norm a" by (simp add: eq assms [symmetric]) finally show "f a = b" @@ -1453,7 +1514,7 @@ by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute) lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" - by (simp add: det_def of_nat_Suc sign_id) + by (simp add: det_def sign_id) lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" proof - @@ -1563,9 +1624,14 @@ by (metis eq_id_iff matrix_id orthogonal_transformation_id that) next case False - with that show thesis - by (auto simp: eq linear_cmul orthogonal_transformation_def - intro: rotation_exists_1 [of "a /\<^sub>R norm a" "b /\<^sub>R norm b", OF 2]) + then obtain f where f: "orthogonal_transformation f" "det (matrix f) = 1" + and f': "f (a /\<^sub>R norm a) = b /\<^sub>R norm b" + using rotation_exists_1 [of "a /\<^sub>R norm a" "b /\<^sub>R norm b", OF 2] by auto + then interpret linear f by (simp add: orthogonal_transformation) + have "f a = b" + using f' False + by (simp add: eq scale) + with f show thesis .. qed lemma rotation_rightward_line: diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Wed May 02 13:49:38 2018 +0200 @@ -1226,7 +1226,7 @@ have "negligible (span S)" using \a \ 0\ a negligible_hyperplane by (blast intro: negligible_subset) then show ?thesis - using span_inc by (blast intro: negligible_subset) + using span_base by (blast intro: negligible_subset) qed proposition negligible_convex_frontier: @@ -1240,7 +1240,7 @@ and spanB: "S \ span B" and cardB: "card B = dim S" by (metis basis_exists) consider "dim S < DIM('N)" | "dim S = DIM('N)" - using dim_subset_UNIV le_eq_less_or_eq by blast + using dim_subset_UNIV le_eq_less_or_eq by auto then show ?thesis proof cases case 1 @@ -2848,8 +2848,12 @@ show "negligible (frontier (g ` K \ g ` L))" by (simp add: negligible_convex_frontier convex_Int conv convex_linear_image that) next - show "negligible (interior (g ` K \ g ` L))" - by (metis (mono_tags, lifting) True finite.emptyI image_Int image_empty interior_Int interior_injective_linear_image intint negligible_finite pairwiseD that) + have "\p N. pairwise p N = (\Na. (Na::'n set) \ N \ (\Nb. Nb \ N \ Na \ Nb \ p Na Nb))" + by (metis pairwise_def) + then have "interior K \ interior L = {}" + using intint that(2) that(3) that(4) by presburger + then show "negligible (interior (g ` K \ g ` L))" + by (metis True empty_imp_negligible image_Int image_empty interior_Int interior_injective_linear_image that(1)) qed moreover have "g ` K \ g ` L \ frontier (g ` K \ g ` L) \ interior (g ` K \ g ` L)" apply (auto simp: frontier_def) @@ -3176,7 +3180,7 @@ assumes "f absolutely_integrable_on S" shows "(norm o f) absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def o_def) - + lemma absolutely_integrable_abs: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Euclidean_Space.thy --- a/src/HOL/Analysis/Euclidean_Space.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Euclidean_Space.thy Wed May 02 13:49:38 2018 +0200 @@ -152,6 +152,53 @@ using False by (auto simp: inner_sum_left) qed +lemma norm_le_componentwise: + "(\b. b \ Basis \ abs(inner x b) \ abs(inner y b)) \ norm x \ norm y" + by (auto simp: norm_le euclidean_inner [of x x] euclidean_inner [of y y] abs_le_square_iff power2_eq_square intro!: sum_mono) + +lemma Basis_le_norm: "b \ Basis \ \inner x b\ \ norm x" + by (rule order_trans [OF Cauchy_Schwarz_ineq2]) simp + +lemma norm_bound_Basis_le: "b \ Basis \ norm x \ e \ \inner x b\ \ e" + by (metis Basis_le_norm order_trans) + +lemma norm_bound_Basis_lt: "b \ Basis \ norm x < e \ \inner x b\ < e" + by (metis Basis_le_norm le_less_trans) + +lemma norm_le_l1: "norm x \ (\b\Basis. \inner x b\)" + apply (subst euclidean_representation[of x, symmetric]) + apply (rule order_trans[OF norm_sum]) + apply (auto intro!: sum_mono) + done + +lemma sum_norm_allsubsets_bound: + fixes f :: "'a \ 'n::euclidean_space" + assumes fP: "finite P" + and fPs: "\Q. Q \ P \ norm (sum f Q) \ e" + shows "(\x\P. norm (f x)) \ 2 * real DIM('n) * e" +proof - + have "(\x\P. norm (f x)) \ (\x\P. \b\Basis. \inner (f x) b\)" + by (rule sum_mono) (rule norm_le_l1) + also have "(\x\P. \b\Basis. \inner (f x) b\) = (\b\Basis. \x\P. \inner (f x) b\)" + by (rule sum.swap) + also have "\ \ of_nat (card (Basis :: 'n set)) * (2 * e)" + proof (rule sum_bounded_above) + fix i :: 'n + assume i: "i \ Basis" + have "norm (\x\P. \inner (f x) i\) \ + norm (inner (\x\P \ - {x. inner (f x) i < 0}. f x) i) + norm (inner (\x\P \ {x. inner (f x) i < 0}. f x) i)" + by (simp add: abs_real_def sum.If_cases[OF fP] sum_negf norm_triangle_ineq4 inner_sum_left + del: real_norm_def) + also have "\ \ e + e" + unfolding real_norm_def + by (intro add_mono norm_bound_Basis_le i fPs) auto + finally show "(\x\P. \inner (f x) i\) \ 2*e" by simp + qed + also have "\ = 2 * real DIM('n) * e" by simp + finally show ?thesis . +qed + + subsection%unimportant \Subclass relationships\ instance euclidean_space \ perfect_space @@ -250,4 +297,56 @@ end + +subsection \Locale instances\ + +lemma finite_dimensional_vector_space_euclidean: + "finite_dimensional_vector_space ( *\<^sub>R) Basis" +proof unfold_locales + show "finite (Basis::'a set)" by (metis finite_Basis) + show "real_vector.independent (Basis::'a set)" + unfolding dependent_def dependent_raw_def[symmetric] + apply (subst span_finite) + apply simp + apply clarify + apply (drule_tac f="inner a" in arg_cong) + apply (simp add: inner_Basis inner_sum_right eq_commute) + done + show "module.span ( *\<^sub>R) Basis = UNIV" + unfolding span_finite [OF finite_Basis] span_raw_def[symmetric] + by (auto intro!: euclidean_representation[symmetric]) +qed + +interpretation eucl?: finite_dimensional_vector_space "scaleR :: real => 'a => 'a::euclidean_space" "Basis" + rewrites "module.dependent ( *\<^sub>R) = dependent" + and "module.representation ( *\<^sub>R) = representation" + and "module.subspace ( *\<^sub>R) = subspace" + and "module.span ( *\<^sub>R) = span" + and "vector_space.extend_basis ( *\<^sub>R) = extend_basis" + and "vector_space.dim ( *\<^sub>R) = dim" + and "Vector_Spaces.linear ( *\<^sub>R) ( *\<^sub>R) = linear" + and "Vector_Spaces.linear ( * ) ( *\<^sub>R) = linear" + and "finite_dimensional_vector_space.dimension Basis = DIM('a)" + and "dimension = DIM('a)" + by (auto simp add: dependent_raw_def representation_raw_def + subspace_raw_def span_raw_def extend_basis_raw_def dim_raw_def linear_def + real_scaleR_def[abs_def] + finite_dimensional_vector_space.dimension_def + intro!: finite_dimensional_vector_space.dimension_def + finite_dimensional_vector_space_euclidean) + +interpretation eucl?: finite_dimensional_vector_space_prod scaleR scaleR Basis Basis + rewrites "Basis_pair = Basis" + and "module_prod.scale ( *\<^sub>R) ( *\<^sub>R) = (scaleR::_\_\('a \ 'b))" +proof - + show "finite_dimensional_vector_space_prod ( *\<^sub>R) ( *\<^sub>R) Basis Basis" + by unfold_locales + interpret finite_dimensional_vector_space_prod "( *\<^sub>R)" "( *\<^sub>R)" "Basis::'a set" "Basis::'b set" + by fact + show "Basis_pair = Basis" + unfolding Basis_pair_def Basis_prod_def by auto + show "module_prod.scale ( *\<^sub>R) ( *\<^sub>R) = scaleR" + by (fact module_prod_scale_eq_scaleR) +qed + end diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Analysis/Finite_Cartesian_Product.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Wed May 02 13:49:38 2018 +0200 @@ -10,7 +10,6 @@ L2_Norm "HOL-Library.Numeral_Type" "HOL-Library.Countable_Set" - "HOL-Library.FuncSet" begin subsection \Finite Cartesian products, with indexing and lambdas\ @@ -230,6 +229,80 @@ by standard (simp_all add: vec_eq_iff) +subsection\Basic componentwise operations on vectors\ + +instantiation vec :: (times, finite) times +begin + +definition "( * ) \ (\ x y. (\ i. (x$i) * (y$i)))" +instance .. + +end + +instantiation vec :: (one, finite) one +begin + +definition "1 \ (\ i. 1)" +instance .. + +end + +instantiation vec :: (ord, finite) ord +begin + +definition "x \ y \ (\i. x$i \ y$i)" +definition "x < (y::'a^'b) \ x \ y \ \ y \ x" +instance .. + +end + +text\The ordering on one-dimensional vectors is linear.\ + +class cart_one = + assumes UNIV_one: "card (UNIV :: 'a set) = Suc 0" +begin + +subclass finite +proof + from UNIV_one show "finite (UNIV :: 'a set)" + by (auto intro!: card_ge_0_finite) +qed + +end + +instance vec:: (order, finite) order + by standard (auto simp: less_eq_vec_def less_vec_def vec_eq_iff + intro: order.trans order.antisym order.strict_implies_order) + +instance vec :: (linorder, cart_one) linorder +proof + obtain a :: 'b where all: "\P. (\i. P i) \ P a" + proof - + have "card (UNIV :: 'b set) = Suc 0" by (rule UNIV_one) + then obtain b :: 'b where "UNIV = {b}" by (auto iff: card_Suc_eq) + then have "\P. (\i\UNIV. P i) \ P b" by auto + then show thesis by (auto intro: that) + qed + fix x y :: "'a^'b::cart_one" + note [simp] = less_eq_vec_def less_vec_def all vec_eq_iff field_simps + show "x \ y \ y \ x" by auto +qed + +text\Constant Vectors\ + +definition "vec x = (\ i. x)" + +text\Also the scalar-vector multiplication.\ + +definition vector_scalar_mult:: "'a::times \ 'a ^ 'n \ 'a ^ 'n" (infixl "*s" 70) + where "c *s x = (\ i. c * (x$i))" + +text \scalar product\ + +definition scalar_product :: "'a :: semiring_1 ^ 'n \ 'a ^ 'n \ 'a" where + "scalar_product v w = (\ i \ UNIV. v $ i * w $ i)" + + subsection \Real vector space\ instantiation vec :: (real_vector, finite) real_vector @@ -543,6 +616,29 @@ unfolding norm_vec_def by (rule member_le_L2_set) simp_all +lemma norm_le_componentwise_cart: + fixes x :: "'a::real_normed_vector^'n" + assumes "\i. norm(x$i) \ norm(y$i)" + shows "norm x \ norm y" + unfolding norm_vec_def + by (rule L2_set_mono) (auto simp: assms) + +lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero) + +lemma component_le_norm_cart: "\x$i\ \ norm x" + apply (simp add: norm_vec_def) + apply (rule member_le_L2_set, simp_all) + done + +lemma norm_bound_component_le_cart: "norm x \ e ==> \x$i\ \ e" + by (metis component_le_norm_cart order_trans) + +lemma norm_bound_component_lt_cart: "norm x < e ==> \x$i\ < e" + by (metis component_le_norm_cart le_less_trans) + +lemma norm_le_l1_cart: "norm x \ sum(\i. \x$i\) UNIV" + by (simp add: norm_vec_def L2_set_le_sum) + lemma bounded_linear_vec_nth: "bounded_linear (\x. x $ i)" apply standard apply (rule vector_add_component) @@ -649,6 +745,16 @@ end +lemma norm_axis_1 [simp]: "norm (axis m (1::real)) = 1" + by (simp add: inner_axis' norm_eq_1) + +lemma sum_norm_allsubsets_bound_cart: + fixes f:: "'a \ real ^'n" + assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (sum f Q) \ e" + shows "sum (\x. norm (f x)) P \ 2 * real CARD('n) * e" + using sum_norm_allsubsets_bound[OF assms] + by simp + lemma cart_eq_inner_axis: "a $ i = inner a (axis i 1)" by (simp add: inner_axis) @@ -684,4 +790,414 @@ shows "(axis_index (axis u 1 :: real^'n)) = (u::'n)" by (simp add: axis_eq_axis axis_index_def) +subsection \A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\ + +lemma sum_cong_aux: + "(\x. x \ A \ f x = g x) \ sum f A = sum g A" + by (auto intro: sum.cong) + +hide_fact (open) sum_cong_aux + +method_setup vector = \ +let + val ss1 = + simpset_of (put_simpset HOL_basic_ss @{context} + addsimps [@{thm sum.distrib} RS sym, + @{thm sum_subtractf} RS sym, @{thm sum_distrib_left}, + @{thm sum_distrib_right}, @{thm sum_negf} RS sym]) + val ss2 = + simpset_of (@{context} addsimps + [@{thm plus_vec_def}, @{thm times_vec_def}, + @{thm minus_vec_def}, @{thm uminus_vec_def}, + @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def}, + @{thm scaleR_vec_def}, + @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}]) + fun vector_arith_tac ctxt ths = + simp_tac (put_simpset ss1 ctxt) + THEN' (fn i => resolve_tac ctxt @{thms Finite_Cartesian_Product.sum_cong_aux} i + ORELSE resolve_tac ctxt @{thms sum.neutral} i + ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i) + (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) + THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths) +in + Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (vector_arith_tac ctxt ths)) end +\ "lift trivial vector statements to real arith statements" + +lemma vec_0[simp]: "vec 0 = 0" by vector +lemma vec_1[simp]: "vec 1 = 1" by vector + +lemma vec_inj[simp]: "vec x = vec y \ x = y" by vector + +lemma vec_in_image_vec: "vec x \ (vec ` S) \ x \ S" by auto + +lemma vec_add: "vec(x + y) = vec x + vec y" by vector +lemma vec_sub: "vec(x - y) = vec x - vec y" by vector +lemma vec_cmul: "vec(c * x) = c *s vec x " by vector +lemma vec_neg: "vec(- x) = - vec x " by vector + +lemma vec_scaleR: "vec(c * x) = c *\<^sub>R vec x" + by vector + +lemma vec_sum: + assumes "finite S" + shows "vec(sum f S) = sum (vec \ f) S" + using assms +proof induct + case empty + then show ?case by simp +next + case insert + then show ?case by (auto simp add: vec_add) +qed + +text\Obvious "component-pushing".\ + +lemma vec_component [simp]: "vec x $ i = x" + by vector + +lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i" + by vector + +lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)" + by vector + +lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector + +lemmas vector_component = + vec_component vector_add_component vector_mult_component + vector_smult_component vector_minus_component vector_uminus_component + vector_scaleR_component cond_component + + +subsection \Some frequently useful arithmetic lemmas over vectors\ + +instance vec :: (semigroup_mult, finite) semigroup_mult + by standard (vector mult.assoc) + +instance vec :: (monoid_mult, finite) monoid_mult + by standard vector+ + +instance vec :: (ab_semigroup_mult, finite) ab_semigroup_mult + by standard (vector mult.commute) + +instance vec :: (comm_monoid_mult, finite) comm_monoid_mult + by standard vector + +instance vec :: (semiring, finite) semiring + by standard (vector field_simps)+ + +instance vec :: (semiring_0, finite) semiring_0 + by standard (vector field_simps)+ +instance vec :: (semiring_1, finite) semiring_1 + by standard vector +instance vec :: (comm_semiring, finite) comm_semiring + by standard (vector field_simps)+ + +instance vec :: (comm_semiring_0, finite) comm_semiring_0 .. +instance vec :: (semiring_0_cancel, finite) semiring_0_cancel .. +instance vec :: (comm_semiring_0_cancel, finite) comm_semiring_0_cancel .. +instance vec :: (ring, finite) ring .. +instance vec :: (semiring_1_cancel, finite) semiring_1_cancel .. +instance vec :: (comm_semiring_1, finite) comm_semiring_1 .. + +instance vec :: (ring_1, finite) ring_1 .. + +instance vec :: (real_algebra, finite) real_algebra + by standard (simp_all add: vec_eq_iff) + +instance vec :: (real_algebra_1, finite) real_algebra_1 .. + +lemma of_nat_index: "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" +proof (induct n) + case 0 + then show ?case by vector +next + case Suc + then show ?case by vector +qed + +lemma one_index [simp]: "(1 :: 'a :: one ^ 'n) $ i = 1" + by vector + +lemma neg_one_index [simp]: "(- 1 :: 'a :: {one, uminus} ^ 'n) $ i = - 1" + by vector + +instance vec :: (semiring_char_0, finite) semiring_char_0 +proof + fix m n :: nat + show "inj (of_nat :: nat \ 'a ^ 'b)" + by (auto intro!: injI simp add: vec_eq_iff of_nat_index) +qed + +instance vec :: (numeral, finite) numeral .. +instance vec :: (semiring_numeral, finite) semiring_numeral .. + +lemma numeral_index [simp]: "numeral w $ i = numeral w" + by (induct w) (simp_all only: numeral.simps vector_add_component one_index) + +lemma neg_numeral_index [simp]: "- numeral w $ i = - numeral w" + by (simp only: vector_uminus_component numeral_index) + +instance vec :: (comm_ring_1, finite) comm_ring_1 .. +instance vec :: (ring_char_0, finite) ring_char_0 .. + +lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x" + by (vector mult.assoc) +lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" + by (vector field_simps) +lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" + by (vector field_simps) +lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector +lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector +lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y" + by (vector field_simps) +lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector +lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector +lemma vector_sneg_minus1: "-x = (-1::'a::ring_1) *s x" by vector +lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector +lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x" + by (vector field_simps) + +lemma vec_eq[simp]: "(vec m = vec n) \ (m = n)" + by (simp add: vec_eq_iff) + +lemma Vector_Spaces_linear_vec [simp]: "Vector_Spaces.linear ( * ) vector_scalar_mult vec" + by unfold_locales (vector algebra_simps)+ + +lemma vector_mul_eq_0[simp]: "(a *s x = 0) \ a = (0::'a::idom) \ x = 0" + by vector + +lemma vector_mul_lcancel[simp]: "a *s x = a *s y \ a = (0::'a::field) \ x = y" + by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib) + +lemma vector_mul_rcancel[simp]: "a *s x = b *s x \ (a::'a::field) = b \ x = 0" + by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib) + +lemma vector_mul_lcancel_imp: "a \ (0::'a::field) ==> a *s x = a *s y ==> (x = y)" + by (metis vector_mul_lcancel) + +lemma vector_mul_rcancel_imp: "x \ 0 \ (a::'a::field) *s x = b *s x ==> a = b" + by (metis vector_mul_rcancel) + +lemma scalar_mult_eq_scaleR [abs_def]: "c *s x = c *\<^sub>R x" + unfolding scaleR_vec_def vector_scalar_mult_def by simp + +lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \c\ * dist x y" + unfolding dist_norm scalar_mult_eq_scaleR + unfolding scaleR_right_diff_distrib[symmetric] by simp + +lemma sum_component [simp]: + fixes f:: " 'a \ ('b::comm_monoid_add) ^'n" + shows "(sum f S)$i = sum (\x. (f x)$i) S" +proof (cases "finite S") + case True + then show ?thesis by induct simp_all +next + case False + then show ?thesis by simp +qed + +lemma sum_eq: "sum f S = (\ i. sum (\x. (f x)$i ) S)" + by (simp add: vec_eq_iff) + +lemma sum_cmul: + fixes f:: "'c \ ('a::semiring_1)^'n" + shows "sum (\x. c *s f x) S = c *s sum f S" + by (simp add: vec_eq_iff sum_distrib_left) + +lemma linear_vec [simp]: "linear vec" + using Vector_Spaces_linear_vec + apply (auto ) + by unfold_locales (vector algebra_simps)+ + +subsection \Matrix operations\ + +text\Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"}\ + +definition map_matrix::"('a \ 'b) \ (('a, 'i::finite)vec, 'j::finite) vec \ (('b, 'i)vec, 'j) vec" where + "map_matrix f x = (\ i j. f (x $ i $ j))" + +lemma nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)" + by (simp add: map_matrix_def) + +definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" + (infixl "**" 70) + where "m ** m' == (\ i j. sum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" + +definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" + (infixl "*v" 70) + where "m *v x \ (\ i. sum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" + +definition vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " + (infixl "v*" 70) + where "v v* m == (\ j. sum (\i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" + +definition "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" +definition transpose where + "(transpose::'a^'n^'m \ 'a^'m^'n) A = (\ i j. ((A$j)$i))" +definition "(row::'m => 'a ^'n^'m \ 'a ^'n) i A = (\ j. ((A$i)$j))" +definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\ i. ((A$i)$j))" +definition "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" +definition "columns(A::'a^'n^'m) = { column i A | i. i \ (UNIV :: 'n set)}" + +lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def) +lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)" + by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps) + +lemma matrix_mul_lid [simp]: + fixes A :: "'a::semiring_1 ^ 'm ^ 'n" + shows "mat 1 ** A = A" + apply (simp add: matrix_matrix_mult_def mat_def) + apply vector + apply (auto simp only: if_distrib if_distribR sum.delta'[OF finite] + mult_1_left mult_zero_left if_True UNIV_I) + done + +lemma matrix_mul_rid [simp]: + fixes A :: "'a::semiring_1 ^ 'm ^ 'n" + shows "A ** mat 1 = A" + apply (simp add: matrix_matrix_mult_def mat_def) + apply vector + apply (auto simp only: if_distrib if_distribR sum.delta[OF finite] + mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) + done + +lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" + apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc) + apply (subst sum.swap) + apply simp + done + +lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" + apply (vector matrix_matrix_mult_def matrix_vector_mult_def + sum_distrib_left sum_distrib_right mult.assoc) + apply (subst sum.swap) + apply simp + done + +lemma matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)" + apply (vector matrix_vector_mult_def mat_def) + apply (simp add: if_distrib if_distribR sum.delta' cong del: if_weak_cong) + done + +lemma matrix_transpose_mul: + "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" + by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute) + +lemma matrix_eq: + fixes A B :: "'a::semiring_1 ^ 'n ^ 'm" + shows "A = B \ (\x. A *v x = B *v x)" (is "?lhs \ ?rhs") + apply auto + apply (subst vec_eq_iff) + apply clarify + apply (clarsimp simp add: matrix_vector_mult_def if_distrib if_distribR vec_eq_iff cong del: if_weak_cong) + apply (erule_tac x="axis ia 1" in allE) + apply (erule_tac x="i" in allE) + apply (auto simp add: if_distrib if_distribR axis_def + sum.delta[OF finite] cong del: if_weak_cong) + done + +lemma matrix_vector_mul_component: "((A::real^_^_) *v x)$k = inner (A$k) x" + by (simp add: matrix_vector_mult_def inner_vec_def) + +lemma dot_lmul_matrix: "inner ((x::real ^_) v* A) y = inner x (A *v y)" + apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def sum_distrib_right sum_distrib_left ac_simps) + apply (subst sum.swap) + apply simp + done + +lemma transpose_mat [simp]: "transpose (mat n) = mat n" + by (vector transpose_def mat_def) + +lemma transpose_transpose [simp]: "transpose(transpose A) = A" + by (vector transpose_def) + +lemma row_transpose [simp]: "row i (transpose A) = column i A" + by (simp add: row_def column_def transpose_def vec_eq_iff) + +lemma column_transpose [simp]: "column i (transpose A) = row i A" + by (simp add: row_def column_def transpose_def vec_eq_iff) + +lemma rows_transpose [simp]: "rows(transpose A) = columns A" + by (auto simp add: rows_def columns_def intro: set_eqI) + +lemma columns_transpose [simp]: "columns(transpose A) = rows A" + by (metis transpose_transpose rows_transpose) + +lemma matrix_mult_sum: + "(A::'a::comm_semiring_1^'n^'m) *v x = sum (\i. (x$i) *s column i A) (UNIV:: 'n set)" + by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute) + +lemma vector_componentwise: + "(x::'a::ring_1^'n) = (\ j. \i\UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)" + by (simp add: axis_def if_distrib sum.If_cases vec_eq_iff) + +lemma basis_expansion: "sum (\i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)" + by (auto simp add: axis_def vec_eq_iff if_distrib sum.If_cases cong del: if_weak_cong) + + +text\Correspondence between matrices and linear operators.\ + +definition matrix :: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" + where "matrix f = (\ i j. (f(axis j 1))$i)" + +lemma matrix_id_mat_1: "matrix id = mat 1" + by (simp add: mat_def matrix_def axis_def) + +lemma matrix_scaleR: "(matrix (( *\<^sub>R) r)) = mat r" + by (simp add: mat_def matrix_def axis_def if_distrib cong: if_cong) + +lemma matrix_vector_mul_linear[intro, simp]: "linear (\x. A *v (x::real ^ _))" + by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff + field_simps sum_distrib_left sum.distrib) + +lemma matrix_vector_mult_add_distrib [algebra_simps]: + "A *v (x + y) = A *v x + A *v y" + by (vector matrix_vector_mult_def sum.distrib distrib_left) + +lemma matrix_vector_mult_diff_distrib [algebra_simps]: + fixes A :: "'a::ring_1^'n^'m" + shows "A *v (x - y) = A *v x - A *v y" + by (vector matrix_vector_mult_def sum_subtractf right_diff_distrib) + +lemma matrix_vector_mult_scaleR[algebra_simps]: + fixes A :: "real^'n^'m" + shows "A *v (c *\<^sub>R x) = c *\<^sub>R (A *v x)" + using linear_iff matrix_vector_mul_linear by blast + +lemma matrix_vector_mult_0_right [simp]: "A *v 0 = 0" + by (simp add: matrix_vector_mult_def vec_eq_iff) + +lemma matrix_vector_mult_0 [simp]: "0 *v w = 0" + by (simp add: matrix_vector_mult_def vec_eq_iff) + +lemma matrix_vector_mult_add_rdistrib [algebra_simps]: + "(A + B) *v x = (A *v x) + (B *v x)" + by (vector matrix_vector_mult_def sum.distrib distrib_right) + +lemma matrix_vector_mult_diff_rdistrib [algebra_simps]: + fixes A :: "'a :: ring_1^'n^'m" + shows "(A - B) *v x = (A *v x) - (B *v x)" + by (vector matrix_vector_mult_def sum_subtractf left_diff_distrib) + +lemma matrix_vector_column: + "(A::'a::comm_semiring_1^'n^_) *v x = sum (\i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" + by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute) + +subsection\Inverse matrices (not necessarily square)\ + +definition + "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" + +definition + "matrix_inv(A:: 'a::semiring_1^'n^'m) = + (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" + +lemma inj_matrix_vector_mult: + fixes A::"'a::field^'n^'m" + assumes "invertible A" + shows "inj (( *v) A)" + by (metis assms inj_on_inverseI invertible_def matrix_vector_mul_assoc matrix_vector_mul_lid) + +end diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Further_Topology.thy Wed May 02 13:49:38 2018 +0200 @@ -68,7 +68,7 @@ by (simp add: subspace_orthogonal_to_vectors T'_def) have dim_eq: "dim T' + dim T = DIM('a)" using dim_subspace_orthogonal_to_vectors [of T UNIV] \subspace T\ - by (simp add: dim_UNIV T'_def) + by (simp add: T'_def) have "\v1 v2. v1 \ span T \ (\w \ span T. orthogonal v2 w) \ x = v1 + v2" for x by (force intro: orthogonal_subspace_decomp_exists [of T x]) then obtain p1 p2 where p1span: "p1 x \ span T" @@ -76,25 +76,35 @@ and eq: "p1 x + p2 x = x" for x by metis then have p1: "\z. p1 z \ T" and ortho: "\w. w \ T \ orthogonal (p2 x) w" for x - using span_eq \subspace T\ by blast+ + using span_eq_iff \subspace T\ by blast+ then have p2: "\z. p2 z \ T'" by (simp add: T'_def orthogonal_commute) have p12_eq: "\x y. \x \ T; y \ T'\ \ p1(x + y) = x \ p2(x + y) = y" proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T']) show "\x y. \x \ T; y \ T'\ \ p2 (x + y) \ span T'" - using span_eq p2 \subspace T'\ by blast + using span_eq_iff p2 \subspace T'\ by blast show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast - qed (auto simp: span_superset) + qed (auto simp: span_base) then have "\c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" - by (metis eq \subspace T\ \subspace T'\ p1 p2 scaleR_add_right subspace_mul) + proof - + fix c :: real and x :: 'a + have f1: "c *\<^sub>R x = c *\<^sub>R p1 x + c *\<^sub>R p2 x" + by (metis eq pth_6) + have f2: "c *\<^sub>R p2 x \ T'" + by (simp add: \subspace T'\ p2 subspace_scale) + have "c *\<^sub>R p1 x \ T" + by (metis (full_types) assms(2) p1span span_eq_iff subspace_scale) + then show "p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" + using f2 f1 p12_eq by presburger + qed moreover have lin_add: "\x y. p1 (x + y) = p1 x + p1 y \ p2 (x + y) = p2 x + p2 y" proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T']) show "\x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)" by (simp add: add.assoc add.left_commute eq) show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast - qed (auto simp: p1span p2 span_superset subspace_add) + qed (auto simp: p1span p2 span_base span_add) ultimately have "linear p1" "linear p2" by unfold_locales auto have "(\z. g (p1 z)) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" @@ -263,7 +273,7 @@ show ?thesis proof (rule that [OF \subspace T\]) show "T \ U" - using span_eq \subspace U\ \T \ span U\ by blast + using span_eq_iff \subspace U\ \T \ span U\ by blast show "aff_dim T = aff_dim S" using dimT \subspace T\ affS aff_dim_subspace by fastforce show "rel_frontier S homeomorphic sphere 0 1 \ T" @@ -314,7 +324,7 @@ where "subspace T'" and affT': "aff_dim T' = aff_dim T" and homT: "rel_frontier T homeomorphic sphere 0 1 \ T'" apply (rule spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a]) - apply (simp add: dim_UNIV aff_dim_le_DIM) + apply (simp add: aff_dim_le_DIM) using \T \ {}\ by blast with homeomorphic_imp_homotopy_eqv have relT: "sphere 0 1 \ T' homotopy_eqv rel_frontier T" @@ -2003,7 +2013,7 @@ have "U \ S" using ope openin_imp_subset by blast have "(UNIV::'b set) homeomorphic S" - by (simp add: \subspace S\ dimS dim_UNIV homeomorphic_subspaces) + by (simp add: \subspace S\ dimS homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k" using homeomorphic_def by blast have homkh: "homeomorphism S (k ` S) k h" @@ -2090,7 +2100,7 @@ proof - obtain V' where "subspace V'" "V' \ U" "dim V' = dim V" using choose_subspace_of_subspace [OF VU] - by (metis span_eq \subspace U\) + by (metis span_eq_iff \subspace U\) then have "V homeomorphic V'" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V V' h k" @@ -2136,7 +2146,7 @@ proof - obtain T where "subspace T" "T \ U" "dim T = dim V" using choose_subspace_of_subspace [of "dim V" U] - by (metis span_eq \subspace U\ \dim V < dim U\ linear not_le) + by (metis \dim V < dim U\ assms(2) order.strict_implies_order span_eq_iff) then have "V homeomorphic T" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V T h k" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Wed May 02 13:49:38 2018 +0200 @@ -2866,7 +2866,7 @@ then have "?sum b = f b" using Suc_pred'[OF \p > 0\] by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib - cond_application_beta sum.If_cases f0) + if_distribR sum.If_cases f0) also have "{..x. p - x - 1) ` {.. x = 0}"]) apply (simp add: dim_hyperplane [OF \i \ 0\]) apply (metis DIM_positive Suc_pred dd not_le not_less_eq_eq) - apply (metis span_eq subspace_hyperplane) + apply (metis span_eq_iff subspace_hyperplane) done have "subspace (span ((+) (- a) ` S))" using subspace_span by blast @@ -956,7 +956,7 @@ have hcont: "continuous_on A h" and kcont: "continuous_on B k" for A B using \linear h\ \linear k\ linear_continuous_on linear_conv_bounded_linear by blast+ have ihhhh[simp]: "\x. x \ S \ i \ h (x - a) = 0" - using Tsub [THEN subsetD] heq span_inc by fastforce + using Tsub [THEN subsetD] heq span_superset by fastforce have "sphere 0 1 - {i} homeomorphic {x. i \ x = 0}" apply (rule homeomorphic_punctured_sphere_affine) using i diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Inner_Product.thy --- a/src/HOL/Analysis/Inner_Product.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Inner_Product.thy Wed May 02 13:49:38 2018 +0200 @@ -161,6 +161,44 @@ end +lemma square_bound_lemma: + fixes x :: real + shows "x < (1 + x) * (1 + x)" +proof - + have "(x + 1/2)\<^sup>2 + 3/4 > 0" + using zero_le_power2[of "x+1/2"] by arith + then show ?thesis + by (simp add: field_simps power2_eq_square) +qed + +lemma square_continuous: + fixes e :: real + shows "e > 0 \ \d. 0 < d \ (\y. \y - x\ < d \ \y * y - x * x\ < e)" + using isCont_power[OF continuous_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2] + by (force simp add: power2_eq_square) + +lemma norm_eq_0_dot: "norm x = 0 \ inner x x = (0::real)" + by simp (* TODO: delete *) + +lemma norm_triangle_sub: + fixes x y :: "'a::real_normed_vector" + shows "norm x \ norm y + norm (x - y)" + using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps) + +lemma norm_le: "norm x \ norm y \ inner x x \ inner y y" + by (simp add: norm_eq_sqrt_inner) + +lemma norm_lt: "norm x < norm y \ inner x x < inner y y" + by (simp add: norm_eq_sqrt_inner) + +lemma norm_eq: "norm x = norm y \ inner x x = inner y y" + apply (subst order_eq_iff) + apply (auto simp: norm_le) + done + +lemma norm_eq_1: "norm x = 1 \ inner x x = 1" + by (simp add: norm_eq_sqrt_inner) + lemma inner_divide_left: fixes a :: "'a :: {real_inner,real_div_algebra}" shows "inner (a / of_real m) b = (inner a b) / m" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Linear_Algebra.thy --- a/src/HOL/Analysis/Linear_Algebra.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Linear_Algebra.thy Wed May 02 13:49:38 2018 +0200 @@ -23,8 +23,8 @@ show "f (a + b) = f a + f b" by (rule f.add) show "f (a - b) = f a - f b" by (rule f.diff) show "f 0 = 0" by (rule f.zero) - show "f (- a) = - f a" by (rule f.minus) - show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scaleR) + show "f (- a) = - f a" by (rule f.neg) + show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scale) qed lemma bounded_linearI: @@ -34,1312 +34,6 @@ shows "bounded_linear f" using assms by (rule bounded_linear_intro) (* FIXME: duplicate *) -subsection \A generic notion of "hull" (convex, affine, conic hull and closure).\ - -definition%important hull :: "('a set \ bool) \ 'a set \ 'a set" (infixl "hull" 75) - where "S hull s = \{t. S t \ s \ t}" - -lemma hull_same: "S s \ S hull s = s" - unfolding hull_def by auto - -lemma hull_in: "(\T. Ball T S \ S (\T)) \ S (S hull s)" - unfolding hull_def Ball_def by auto - -lemma hull_eq: "(\T. Ball T S \ S (\T)) \ (S hull s) = s \ S s" - using hull_same[of S s] hull_in[of S s] by metis - -lemma hull_hull [simp]: "S hull (S hull s) = S hull s" - unfolding hull_def by blast - -lemma hull_subset[intro]: "s \ (S hull s)" - unfolding hull_def by blast - -lemma hull_mono: "s \ t \ (S hull s) \ (S hull t)" - unfolding hull_def by blast - -lemma hull_antimono: "\x. S x \ T x \ (T hull s) \ (S hull s)" - unfolding hull_def by blast - -lemma hull_minimal: "s \ t \ S t \ (S hull s) \ t" - unfolding hull_def by blast - -lemma subset_hull: "S t \ S hull s \ t \ s \ t" - unfolding hull_def by blast - -lemma hull_UNIV [simp]: "S hull UNIV = UNIV" - unfolding hull_def by auto - -lemma hull_unique: "s \ t \ S t \ (\t'. s \ t' \ S t' \ t \ t') \ (S hull s = t)" - unfolding hull_def by auto - -lemma hull_induct: "(\x. x\ S \ P x) \ Q {x. P x} \ \x\ Q hull S. P x" - using hull_minimal[of S "{x. P x}" Q] - by (auto simp add: subset_eq) - -lemma hull_inc: "x \ S \ x \ P hull S" - by (metis hull_subset subset_eq) - -lemma hull_Un_subset: "(S hull s) \ (S hull t) \ (S hull (s \ t))" - unfolding Un_subset_iff by (metis hull_mono Un_upper1 Un_upper2) - -lemma hull_Un: - assumes T: "\T. Ball T S \ S (\T)" - shows "S hull (s \ t) = S hull (S hull s \ S hull t)" - apply (rule equalityI) - apply (meson hull_mono hull_subset sup.mono) - by (metis hull_Un_subset hull_hull hull_mono) - -lemma hull_Un_left: "P hull (S \ T) = P hull (P hull S \ T)" - apply (rule equalityI) - apply (simp add: Un_commute hull_mono hull_subset sup.coboundedI2) - by (metis Un_subset_iff hull_hull hull_mono hull_subset) - -lemma hull_Un_right: "P hull (S \ T) = P hull (S \ P hull T)" - by (metis hull_Un_left sup.commute) - -lemma hull_insert: - "P hull (insert a S) = P hull (insert a (P hull S))" - by (metis hull_Un_right insert_is_Un) - -lemma hull_redundant_eq: "a \ (S hull s) \ S hull (insert a s) = S hull s" - unfolding hull_def by blast - -lemma hull_redundant: "a \ (S hull s) \ S hull (insert a s) = S hull s" - by (metis hull_redundant_eq) - -subsection \Linear functions.\ - -lemma%important linear_iff: - "linear f \ (\x y. f (x + y) = f x + f y) \ (\c x. f (c *\<^sub>R x) = c *\<^sub>R f x)" - (is "linear f \ ?rhs") -proof%unimportant - assume "linear f" - then interpret f: linear f . - show "?rhs" by (simp add: f.add f.scaleR) -next - assume "?rhs" - then show "linear f" by unfold_locales simp_all -qed - -lemma linear_compose_cmul: "linear f \ linear (\x. c *\<^sub>R f x)" - by (simp add: linear_iff algebra_simps) - -lemma linear_compose_scaleR: "linear f \ linear (\x. f x *\<^sub>R c)" - by (simp add: linear_iff scaleR_add_left) - -lemma linear_compose_neg: "linear f \ linear (\x. - f x)" - by (simp add: linear_iff) - -lemma linear_compose_add: "linear f \ linear g \ linear (\x. f x + g x)" - by (simp add: linear_iff algebra_simps) - -lemma linear_compose_sub: "linear f \ linear g \ linear (\x. f x - g x)" - by (simp add: linear_iff algebra_simps) - -lemma linear_compose: "linear f \ linear g \ linear (g \ f)" - by (simp add: linear_iff) - -lemma linear_id: "linear id" - by (simp add: linear_iff id_def) - -lemma linear_zero: "linear (\x. 0)" - by (simp add: linear_iff) - -lemma linear_uminus: "linear uminus" -by (simp add: linear_iff) - -lemma linear_compose_sum: - assumes lS: "\a \ S. linear (f a)" - shows "linear (\x. sum (\a. f a x) S)" -proof (cases "finite S") - case True - then show ?thesis - using lS by induct (simp_all add: linear_zero linear_compose_add) -next - case False - then show ?thesis - by (simp add: linear_zero) -qed - -lemma linear_0: "linear f \ f 0 = 0" - unfolding linear_iff - apply clarsimp - apply (erule allE[where x="0::'a"]) - apply simp - done - -lemma linear_cmul: "linear f \ f (c *\<^sub>R x) = c *\<^sub>R f x" - by (rule linear.scaleR) - -lemma linear_neg: "linear f \ f (- x) = - f x" - using linear_cmul [where c="-1"] by simp - -lemma linear_add: "linear f \ f (x + y) = f x + f y" - by (metis linear_iff) - -lemma linear_diff: "linear f \ f (x - y) = f x - f y" - using linear_add [of f x "- y"] by (simp add: linear_neg) - -lemma linear_sum: - assumes f: "linear f" - shows "f (sum g S) = sum (f \ g) S" -proof (cases "finite S") - case True - then show ?thesis - by induct (simp_all add: linear_0 [OF f] linear_add [OF f]) -next - case False - then show ?thesis - by (simp add: linear_0 [OF f]) -qed - -lemma linear_sum_mul: - assumes lin: "linear f" - shows "f (sum (\i. c i *\<^sub>R v i) S) = sum (\i. c i *\<^sub>R f (v i)) S" - using linear_sum[OF lin, of "\i. c i *\<^sub>R v i" , unfolded o_def] linear_cmul[OF lin] - by simp - -lemma linear_injective_0: - assumes lin: "linear f" - shows "inj f \ (\x. f x = 0 \ x = 0)" -proof - - have "inj f \ (\ x y. f x = f y \ x = y)" - by (simp add: inj_on_def) - also have "\ \ (\ x y. f x - f y = 0 \ x - y = 0)" - by simp - also have "\ \ (\ x y. f (x - y) = 0 \ x - y = 0)" - by (simp add: linear_diff[OF lin]) - also have "\ \ (\ x. f x = 0 \ x = 0)" - by auto - finally show ?thesis . -qed - -lemma linear_scaleR [simp]: "linear (\x. scaleR c x)" - by (simp add: linear_iff scaleR_add_right) - -lemma linear_scaleR_left [simp]: "linear (\r. scaleR r x)" - by (simp add: linear_iff scaleR_add_left) - -lemma injective_scaleR: "c \ 0 \ inj (\x::'a::real_vector. scaleR c x)" - by (simp add: inj_on_def) - -lemma linear_add_cmul: - assumes "linear f" - shows "f (a *\<^sub>R x + b *\<^sub>R y) = a *\<^sub>R f x + b *\<^sub>R f y" - using linear_add[of f] linear_cmul[of f] assms by simp - -subsection \Subspaces of vector spaces\ - -definition%important (in real_vector) subspace :: "'a set \ bool" - where "subspace S \ 0 \ S \ (\x \ S. \y \ S. x + y \ S) \ (\c. \x \ S. c *\<^sub>R x \ S)" - -definition%important (in real_vector) "span S = (subspace hull S)" -definition%important (in real_vector) "dependent S \ (\a \ S. a \ span (S - {a}))" -abbreviation (in real_vector) "independent s \ \ dependent s" - -text \Closure properties of subspaces.\ - -lemma subspace_UNIV[simp]: "subspace UNIV" - by (simp add: subspace_def) - -lemma (in real_vector) subspace_0: "subspace S \ 0 \ S" - by (metis subspace_def) - -lemma (in real_vector) subspace_add: "subspace S \ x \ S \ y \ S \ x + y \ S" - by (metis subspace_def) - -lemma (in real_vector) subspace_mul: "subspace S \ x \ S \ c *\<^sub>R x \ S" - by (metis subspace_def) - -lemma subspace_neg: "subspace S \ x \ S \ - x \ S" - by (metis scaleR_minus1_left subspace_mul) - -lemma subspace_diff: "subspace S \ x \ S \ y \ S \ x - y \ S" - using subspace_add [of S x "- y"] by (simp add: subspace_neg) - -lemma (in real_vector) subspace_sum: - assumes sA: "subspace A" - and f: "\x. x \ B \ f x \ A" - shows "sum f B \ A" -proof (cases "finite B") - case True - then show ?thesis - using f by induct (simp_all add: subspace_0 [OF sA] subspace_add [OF sA]) -qed (simp add: subspace_0 [OF sA]) - -lemma subspace_trivial [iff]: "subspace {0}" - by (simp add: subspace_def) - -lemma (in real_vector) subspace_inter: "subspace A \ subspace B \ subspace (A \ B)" - by (simp add: subspace_def) - -lemma subspace_Times: "subspace A \ subspace B \ subspace (A \ B)" - unfolding subspace_def zero_prod_def by simp - -lemma subspace_sums: "\subspace S; subspace T\ \ subspace {x + y|x y. x \ S \ y \ T}" -apply (simp add: subspace_def) -apply (intro conjI impI allI) - using add.right_neutral apply blast - apply clarify - apply (metis add.assoc add.left_commute) -using scaleR_add_right by blast - -subsection%unimportant \Properties of span\ - -lemma (in real_vector) span_mono: "A \ B \ span A \ span B" - by (metis span_def hull_mono) - -lemma (in real_vector) subspace_span [iff]: "subspace (span S)" - unfolding span_def - apply (rule hull_in) - apply (simp only: subspace_def Inter_iff Int_iff subset_eq) - apply auto - done - -lemma (in real_vector) span_clauses: - "a \ S \ a \ span S" - "0 \ span S" - "x\ span S \ y \ span S \ x + y \ span S" - "x \ span S \ c *\<^sub>R x \ span S" - by (metis span_def hull_subset subset_eq) (metis subspace_span subspace_def)+ - -lemma span_unique: - "S \ T \ subspace T \ (\T'. S \ T' \ subspace T' \ T \ T') \ span S = T" - unfolding span_def by (rule hull_unique) - -lemma span_minimal: "S \ T \ subspace T \ span S \ T" - unfolding span_def by (rule hull_minimal) - -lemma span_UNIV [simp]: "span UNIV = UNIV" - by (intro span_unique) auto - -lemma (in real_vector) span_induct: - assumes x: "x \ span S" - and P: "subspace (Collect P)" - and SP: "\x. x \ S \ P x" - shows "P x" -proof - - from SP have SP': "S \ Collect P" - by (simp add: subset_eq) - from x hull_minimal[where S=subspace, OF SP' P, unfolded span_def[symmetric]] - show ?thesis - using subset_eq by force -qed - -lemma span_empty[simp]: "span {} = {0}" - apply (simp add: span_def) - apply (rule hull_unique) - apply (auto simp add: subspace_def) - done - -lemma (in real_vector) independent_empty [iff]: "independent {}" - by (simp add: dependent_def) - -lemma dependent_single[simp]: "dependent {x} \ x = 0" - unfolding dependent_def by auto - -lemma (in real_vector) independent_mono: "independent A \ B \ A \ independent B" - apply (clarsimp simp add: dependent_def span_mono) - apply (subgoal_tac "span (B - {a}) \ span (A - {a})") - apply force - apply (rule span_mono) - apply auto - done - -lemma (in real_vector) span_subspace: "A \ B \ B \ span A \ subspace B \ span A = B" - by (metis order_antisym span_def hull_minimal) - -lemma (in real_vector) span_induct': - "\x \ S. P x \ subspace {x. P x} \ \x \ span S. P x" - unfolding span_def by (rule hull_induct) auto - -inductive_set (in real_vector) span_induct_alt_help for S :: "'a set" -where - span_induct_alt_help_0: "0 \ span_induct_alt_help S" -| span_induct_alt_help_S: - "x \ S \ z \ span_induct_alt_help S \ - (c *\<^sub>R x + z) \ span_induct_alt_help S" - -lemma span_induct_alt': - assumes h0: "h 0" - and hS: "\c x y. x \ S \ h y \ h (c *\<^sub>R x + y)" - shows "\x \ span S. h x" -proof - - { - fix x :: 'a - assume x: "x \ span_induct_alt_help S" - have "h x" - apply (rule span_induct_alt_help.induct[OF x]) - apply (rule h0) - apply (rule hS) - apply assumption - apply assumption - done - } - note th0 = this - { - fix x - assume x: "x \ span S" - have "x \ span_induct_alt_help S" - proof (rule span_induct[where x=x and S=S]) - show "x \ span S" by (rule x) - next - fix x - assume xS: "x \ S" - from span_induct_alt_help_S[OF xS span_induct_alt_help_0, of 1] - show "x \ span_induct_alt_help S" - by simp - next - have "0 \ span_induct_alt_help S" by (rule span_induct_alt_help_0) - moreover - { - fix x y - assume h: "x \ span_induct_alt_help S" "y \ span_induct_alt_help S" - from h have "(x + y) \ span_induct_alt_help S" - apply (induct rule: span_induct_alt_help.induct) - apply simp - unfolding add.assoc - apply (rule span_induct_alt_help_S) - apply assumption - apply simp - done - } - moreover - { - fix c x - assume xt: "x \ span_induct_alt_help S" - then have "(c *\<^sub>R x) \ span_induct_alt_help S" - apply (induct rule: span_induct_alt_help.induct) - apply (simp add: span_induct_alt_help_0) - apply (simp add: scaleR_right_distrib) - apply (rule span_induct_alt_help_S) - apply assumption - apply simp - done } - ultimately show "subspace {a. a \ span_induct_alt_help S}" - unfolding subspace_def Ball_def by blast - qed - } - with th0 show ?thesis by blast -qed - -lemma span_induct_alt: - assumes h0: "h 0" - and hS: "\c x y. x \ S \ h y \ h (c *\<^sub>R x + y)" - and x: "x \ span S" - shows "h x" - using span_induct_alt'[of h S] h0 hS x by blast - -text \Individual closure properties.\ - -lemma span_span: "span (span A) = span A" - unfolding span_def hull_hull .. - -lemma (in real_vector) span_superset: "x \ S \ x \ span S" - by (metis span_clauses(1)) - -lemma (in real_vector) span_0 [simp]: "0 \ span S" - by (metis subspace_span subspace_0) - -lemma span_inc: "S \ span S" - by (metis subset_eq span_superset) - -lemma span_eq: "span S = span T \ S \ span T \ T \ span S" - using span_inc[unfolded subset_eq] using span_mono[of T "span S"] span_mono[of S "span T"] - by (auto simp add: span_span) - -lemma (in real_vector) dependent_0: - assumes "0 \ A" - shows "dependent A" - unfolding dependent_def - using assms span_0 - by blast - -lemma (in real_vector) span_add: "x \ span S \ y \ span S \ x + y \ span S" - by (metis subspace_add subspace_span) - -lemma (in real_vector) span_mul: "x \ span S \ c *\<^sub>R x \ span S" - by (metis subspace_span subspace_mul) - -lemma span_neg: "x \ span S \ - x \ span S" - by (metis subspace_neg subspace_span) - -lemma span_diff: "x \ span S \ y \ span S \ x - y \ span S" - by (metis subspace_span subspace_diff) - -lemma (in real_vector) span_sum: "(\x. x \ A \ f x \ span S) \ sum f A \ span S" - by (rule subspace_sum [OF subspace_span]) - -lemma span_add_eq: "x \ span S \ x + y \ span S \ y \ span S" - by (metis add_minus_cancel scaleR_minus1_left subspace_def subspace_span) - -text \The key breakdown property.\ - -lemma span_singleton: "span {x} = range (\k. k *\<^sub>R x)" -proof (rule span_unique) - show "{x} \ range (\k. k *\<^sub>R x)" - by (fast intro: scaleR_one [symmetric]) - show "subspace (range (\k. k *\<^sub>R x))" - unfolding subspace_def - by (auto intro: scaleR_add_left [symmetric]) -next - fix T - assume "{x} \ T" and "subspace T" - then show "range (\k. k *\<^sub>R x) \ T" - unfolding subspace_def by auto -qed - -text \Mapping under linear image.\ - -lemma subspace_linear_image: - assumes lf: "linear f" - and sS: "subspace S" - shows "subspace (f ` S)" - using lf sS linear_0[OF lf] - unfolding linear_iff subspace_def - apply (auto simp add: image_iff) - apply (rule_tac x="x + y" in bexI) - apply auto - apply (rule_tac x="c *\<^sub>R x" in bexI) - apply auto - done - -lemma subspace_linear_vimage: "linear f \ subspace S \ subspace (f -` S)" - by (auto simp add: subspace_def linear_iff linear_0[of f]) - -lemma subspace_linear_preimage: "linear f \ subspace S \ subspace {x. f x \ S}" - by (auto simp add: subspace_def linear_iff linear_0[of f]) - -lemma span_linear_image: - assumes lf: "linear f" - shows "span (f ` S) = f ` span S" -proof (rule span_unique) - show "f ` S \ f ` span S" - by (intro image_mono span_inc) - show "subspace (f ` span S)" - using lf subspace_span by (rule subspace_linear_image) -next - fix T - assume "f ` S \ T" and "subspace T" - then show "f ` span S \ T" - unfolding image_subset_iff_subset_vimage - by (intro span_minimal subspace_linear_vimage lf) -qed - -lemma spans_image: - assumes lf: "linear f" - and VB: "V \ span B" - shows "f ` V \ span (f ` B)" - unfolding span_linear_image[OF lf] by (metis VB image_mono) - -lemma span_Un: "span (A \ B) = (\(a, b). a + b) ` (span A \ span B)" -proof (rule span_unique) - show "A \ B \ (\(a, b). a + b) ` (span A \ span B)" - by safe (force intro: span_clauses)+ -next - have "linear (\(a, b). a + b)" - by (simp add: linear_iff scaleR_add_right) - moreover have "subspace (span A \ span B)" - by (intro subspace_Times subspace_span) - ultimately show "subspace ((\(a, b). a + b) ` (span A \ span B))" - by (rule subspace_linear_image) -next - fix T - assume "A \ B \ T" and "subspace T" - then show "(\(a, b). a + b) ` (span A \ span B) \ T" - by (auto intro!: subspace_add elim: span_induct) -qed - -lemma span_insert: "span (insert a S) = {x. \k. (x - k *\<^sub>R a) \ span S}" -proof - - have "span ({a} \ S) = {x. \k. (x - k *\<^sub>R a) \ span S}" - unfolding span_Un span_singleton - apply safe - apply (rule_tac x=k in exI, simp) - apply (erule rev_image_eqI [OF SigmaI [OF rangeI]]) - apply auto - done - then show ?thesis by simp -qed - -lemma span_breakdown: - assumes bS: "b \ S" - and aS: "a \ span S" - shows "\k. a - k *\<^sub>R b \ span (S - {b})" - using assms span_insert [of b "S - {b}"] - by (simp add: insert_absorb) - -lemma span_breakdown_eq: "x \ span (insert a S) \ (\k. x - k *\<^sub>R a \ span S)" - by (simp add: span_insert) - -text \Hence some "reversal" results.\ - -lemma in_span_insert: - assumes a: "a \ span (insert b S)" - and na: "a \ span S" - shows "b \ span (insert a S)" -proof - - from a obtain k where k: "a - k *\<^sub>R b \ span S" - unfolding span_insert by fast - show ?thesis - proof (cases "k = 0") - case True - with k have "a \ span S" by simp - with na show ?thesis by simp - next - case False - from k have "(- inverse k) *\<^sub>R (a - k *\<^sub>R b) \ span S" - by (rule span_mul) - then have "b - inverse k *\<^sub>R a \ span S" - using \k \ 0\ by (simp add: scaleR_diff_right) - then show ?thesis - unfolding span_insert by fast - qed -qed - -lemma in_span_delete: - assumes a: "a \ span S" - and na: "a \ span (S - {b})" - shows "b \ span (insert a (S - {b}))" - apply (rule in_span_insert) - apply (rule set_rev_mp) - apply (rule a) - apply (rule span_mono) - apply blast - apply (rule na) - done - -text \Transitivity property.\ - -lemma span_redundant: "x \ span S \ span (insert x S) = span S" - unfolding span_def by (rule hull_redundant) - -lemma span_trans: - assumes x: "x \ span S" - and y: "y \ span (insert x S)" - shows "y \ span S" - using assms by (simp only: span_redundant) - -lemma span_insert_0[simp]: "span (insert 0 S) = span S" - by (simp only: span_redundant span_0) - -text \An explicit expansion is sometimes needed.\ - -lemma span_explicit: - "span P = {y. \S u. finite S \ S \ P \ sum (\v. u v *\<^sub>R v) S = y}" - (is "_ = ?E" is "_ = {y. ?h y}" is "_ = {y. \S u. ?Q S u y}") -proof - - { - fix x - assume "?h x" - then obtain S u where "finite S" and "S \ P" and "sum (\v. u v *\<^sub>R v) S = x" - by blast - then have "x \ span P" - by (auto intro: span_sum span_mul span_superset) - } - moreover - have "\x \ span P. ?h x" - proof (rule span_induct_alt') - show "?h 0" - by (rule exI[where x="{}"], simp) - next - fix c x y - assume x: "x \ P" - assume hy: "?h y" - from hy obtain S u where fS: "finite S" and SP: "S\P" - and u: "sum (\v. u v *\<^sub>R v) S = y" by blast - let ?S = "insert x S" - let ?u = "\y. if y = x then (if x \ S then u y + c else c) else u y" - from fS SP x have th0: "finite (insert x S)" "insert x S \ P" - by blast+ - have "?Q ?S ?u (c*\<^sub>R x + y)" - proof cases - assume xS: "x \ S" - have "sum (\v. ?u v *\<^sub>R v) ?S = (\v\S - {x}. u v *\<^sub>R v) + (u x + c) *\<^sub>R x" - using xS by (simp add: sum.remove [OF fS xS] insert_absorb) - also have "\ = (\v\S. u v *\<^sub>R v) + c *\<^sub>R x" - by (simp add: sum.remove [OF fS xS] algebra_simps) - also have "\ = c*\<^sub>R x + y" - by (simp add: add.commute u) - finally have "sum (\v. ?u v *\<^sub>R v) ?S = c*\<^sub>R x + y" . - then show ?thesis using th0 by blast - next - assume xS: "x \ S" - have th00: "(\v\S. (if v = x then c else u v) *\<^sub>R v) = y" - unfolding u[symmetric] - apply (rule sum.cong) - using xS - apply auto - done - show ?thesis using fS xS th0 - by (simp add: th00 add.commute cong del: if_weak_cong) - qed - then show "?h (c*\<^sub>R x + y)" - by fast - qed - ultimately show ?thesis by blast -qed - -lemma dependent_explicit: - "dependent P \ (\S u. finite S \ S \ P \ (\v\S. u v \ 0 \ sum (\v. u v *\<^sub>R v) S = 0))" - (is "?lhs = ?rhs") -proof - - { - assume dP: "dependent P" - then obtain a S u where aP: "a \ P" and fS: "finite S" - and SP: "S \ P - {a}" and ua: "sum (\v. u v *\<^sub>R v) S = a" - unfolding dependent_def span_explicit by blast - let ?S = "insert a S" - let ?u = "\y. if y = a then - 1 else u y" - let ?v = a - from aP SP have aS: "a \ S" - by blast - from fS SP aP have th0: "finite ?S" "?S \ P" "?v \ ?S" "?u ?v \ 0" - by auto - have s0: "sum (\v. ?u v *\<^sub>R v) ?S = 0" - using fS aS - apply simp - apply (subst (2) ua[symmetric]) - apply (rule sum.cong) - apply auto - done - with th0 have ?rhs by fast - } - moreover - { - fix S u v - assume fS: "finite S" - and SP: "S \ P" - and vS: "v \ S" - and uv: "u v \ 0" - and u: "sum (\v. u v *\<^sub>R v) S = 0" - let ?a = v - let ?S = "S - {v}" - let ?u = "\i. (- u i) / u v" - have th0: "?a \ P" "finite ?S" "?S \ P" - using fS SP vS by auto - have "sum (\v. ?u v *\<^sub>R v) ?S = - sum (\v. (- (inverse (u ?a))) *\<^sub>R (u v *\<^sub>R v)) S - ?u v *\<^sub>R v" - using fS vS uv by (simp add: sum_diff1 field_simps) - also have "\ = ?a" - unfolding scaleR_right.sum [symmetric] u using uv by simp - finally have "sum (\v. ?u v *\<^sub>R v) ?S = ?a" . - with th0 have ?lhs - unfolding dependent_def span_explicit - apply - - apply (rule bexI[where x= "?a"]) - apply (simp_all del: scaleR_minus_left) - apply (rule exI[where x= "?S"]) - apply (auto simp del: scaleR_minus_left) - done - } - ultimately show ?thesis by blast -qed - -lemma dependent_finite: - assumes "finite S" - shows "dependent S \ (\u. (\v \ S. u v \ 0) \ (\v\S. u v *\<^sub>R v) = 0)" - (is "?lhs = ?rhs") -proof - assume ?lhs - then obtain T u v - where "finite T" "T \ S" "v\T" "u v \ 0" "(\v\T. u v *\<^sub>R v) = 0" - by (force simp: dependent_explicit) - with assms show ?rhs - apply (rule_tac x="\v. if v \ T then u v else 0" in exI) - apply (auto simp: sum.mono_neutral_right) - done -next - assume ?rhs with assms show ?lhs - by (fastforce simp add: dependent_explicit) -qed - -lemma span_alt: - "span B = {(\x | f x \ 0. f x *\<^sub>R x) | f. {x. f x \ 0} \ B \ finite {x. f x \ 0}}" - unfolding span_explicit - apply safe - subgoal for x S u - by (intro exI[of _ "\x. if x \ S then u x else 0"]) - (auto intro!: sum.mono_neutral_cong_right) - apply auto - done - -lemma dependent_alt: - "dependent B \ - (\X. finite {x. X x \ 0} \ {x. X x \ 0} \ B \ (\x|X x \ 0. X x *\<^sub>R x) = 0 \ (\x. X x \ 0))" - unfolding dependent_explicit - apply safe - subgoal for S u v - apply (intro exI[of _ "\x. if x \ S then u x else 0"]) - apply (subst sum.mono_neutral_cong_left[where T=S]) - apply (auto intro!: sum.mono_neutral_cong_right cong: rev_conj_cong) - done - apply auto - done - -lemma independent_alt: - "independent B \ - (\X. finite {x. X x \ 0} \ {x. X x \ 0} \ B \ (\x|X x \ 0. X x *\<^sub>R x) = 0 \ (\x. X x = 0))" - unfolding dependent_alt by auto - -lemma independentD_alt: - "independent B \ finite {x. X x \ 0} \ {x. X x \ 0} \ B \ (\x|X x \ 0. X x *\<^sub>R x) = 0 \ X x = 0" - unfolding independent_alt by blast - -lemma independentD_unique: - assumes B: "independent B" - and X: "finite {x. X x \ 0}" "{x. X x \ 0} \ B" - and Y: "finite {x. Y x \ 0}" "{x. Y x \ 0} \ B" - and "(\x | X x \ 0. X x *\<^sub>R x) = (\x| Y x \ 0. Y x *\<^sub>R x)" - shows "X = Y" -proof - - have "X x - Y x = 0" for x - using B - proof (rule independentD_alt) - have "{x. X x - Y x \ 0} \ {x. X x \ 0} \ {x. Y x \ 0}" - by auto - then show "finite {x. X x - Y x \ 0}" "{x. X x - Y x \ 0} \ B" - using X Y by (auto dest: finite_subset) - then have "(\x | X x - Y x \ 0. (X x - Y x) *\<^sub>R x) = (\v\{S. X S \ 0} \ {S. Y S \ 0}. (X v - Y v) *\<^sub>R v)" - using X Y by (intro sum.mono_neutral_cong_left) auto - also have "\ = (\v\{S. X S \ 0} \ {S. Y S \ 0}. X v *\<^sub>R v) - (\v\{S. X S \ 0} \ {S. Y S \ 0}. Y v *\<^sub>R v)" - by (simp add: scaleR_diff_left sum_subtractf assms) - also have "(\v\{S. X S \ 0} \ {S. Y S \ 0}. X v *\<^sub>R v) = (\v\{S. X S \ 0}. X v *\<^sub>R v)" - using X Y by (intro sum.mono_neutral_cong_right) auto - also have "(\v\{S. X S \ 0} \ {S. Y S \ 0}. Y v *\<^sub>R v) = (\v\{S. Y S \ 0}. Y v *\<^sub>R v)" - using X Y by (intro sum.mono_neutral_cong_right) auto - finally show "(\x | X x - Y x \ 0. (X x - Y x) *\<^sub>R x) = 0" - using assms by simp - qed - then show ?thesis - by auto -qed - -text \This is useful for building a basis step-by-step.\ - -lemma independent_insert: - "independent (insert a S) \ - (if a \ S then independent S else independent S \ a \ span S)" - (is "?lhs \ ?rhs") -proof (cases "a \ S") - case True - then show ?thesis - using insert_absorb[OF True] by simp -next - case False - show ?thesis - proof - assume i: ?lhs - then show ?rhs - using False - apply simp - apply (rule conjI) - apply (rule independent_mono) - apply assumption - apply blast - apply (simp add: dependent_def) - done - next - assume i: ?rhs - show ?lhs - using i False - apply (auto simp add: dependent_def) - by (metis in_span_insert insert_Diff_if insert_Diff_single insert_absorb) - qed -qed - -lemma independent_Union_directed: - assumes directed: "\c d. c \ C \ d \ C \ c \ d \ d \ c" - assumes indep: "\c. c \ C \ independent c" - shows "independent (\C)" -proof - assume "dependent (\C)" - then obtain u v S where S: "finite S" "S \ \C" "v \ S" "u v \ 0" "(\v\S. u v *\<^sub>R v) = 0" - by (auto simp: dependent_explicit) - - have "S \ {}" - using \v \ S\ by auto - have "\c\C. S \ c" - using \finite S\ \S \ {}\ \S \ \C\ - proof (induction rule: finite_ne_induct) - case (insert i I) - then obtain c d where cd: "c \ C" "d \ C" and iI: "I \ c" "i \ d" - by blast - from directed[OF cd] cd have "c \ d \ C" - by (auto simp: sup.absorb1 sup.absorb2) - with iI show ?case - by (intro bexI[of _ "c \ d"]) auto - qed auto - then obtain c where "c \ C" "S \ c" - by auto - have "dependent c" - unfolding dependent_explicit - by (intro exI[of _ S] exI[of _ u] bexI[of _ v] conjI) fact+ - with indep[OF \c \ C\] show False - by auto -qed - -text \Hence we can create a maximal independent subset.\ - -lemma maximal_independent_subset_extend: - assumes "S \ V" "independent S" - shows "\B. S \ B \ B \ V \ independent B \ V \ span B" -proof - - let ?C = "{B. S \ B \ independent B \ B \ V}" - have "\M\?C. \X\?C. M \ X \ X = M" - proof (rule subset_Zorn) - fix C :: "'a set set" assume "subset.chain ?C C" - then have C: "\c. c \ C \ c \ V" "\c. c \ C \ S \ c" "\c. c \ C \ independent c" - "\c d. c \ C \ d \ C \ c \ d \ d \ c" - unfolding subset.chain_def by blast+ - - show "\U\?C. \X\C. X \ U" - proof cases - assume "C = {}" with assms show ?thesis - by (auto intro!: exI[of _ S]) - next - assume "C \ {}" - with C(2) have "S \ \C" - by auto - moreover have "independent (\C)" - by (intro independent_Union_directed C) - moreover have "\C \ V" - using C by auto - ultimately show ?thesis - by auto - qed - qed - then obtain B where B: "independent B" "B \ V" "S \ B" - and max: "\S. independent S \ S \ V \ B \ S \ S = B" - by auto - moreover - { assume "\ V \ span B" - then obtain v where "v \ V" "v \ span B" - by auto - with B have "independent (insert v B)" - unfolding independent_insert by auto - from max[OF this] \v \ V\ \B \ V\ - have "v \ B" - by auto - with \v \ span B\ have False - by (auto intro: span_superset) } - ultimately show ?thesis - by (auto intro!: exI[of _ B]) -qed - - -lemma maximal_independent_subset: - "\B. B \ V \ independent B \ V \ span B" - by (metis maximal_independent_subset_extend[of "{}"] empty_subsetI independent_empty) - -lemma span_finite: - assumes fS: "finite S" - shows "span S = {y. \u. sum (\v. u v *\<^sub>R v) S = y}" - (is "_ = ?rhs") -proof - - { - fix y - assume y: "y \ span S" - from y obtain S' u where fS': "finite S'" - and SS': "S' \ S" - and u: "sum (\v. u v *\<^sub>R v) S' = y" - unfolding span_explicit by blast - let ?u = "\x. if x \ S' then u x else 0" - have "sum (\v. ?u v *\<^sub>R v) S = sum (\v. u v *\<^sub>R v) S'" - using SS' fS by (auto intro!: sum.mono_neutral_cong_right) - then have "sum (\v. ?u v *\<^sub>R v) S = y" by (metis u) - then have "y \ ?rhs" by auto - } - moreover - { - fix y u - assume u: "sum (\v. u v *\<^sub>R v) S = y" - then have "y \ span S" using fS unfolding span_explicit by auto - } - ultimately show ?thesis by blast -qed - -lemma linear_independent_extend_subspace: - assumes "independent B" - shows "\g. linear g \ (\x\B. g x = f x) \ range g = span (f`B)" -proof - - from maximal_independent_subset_extend[OF _ \independent B\, of UNIV] - obtain B' where "B \ B'" "independent B'" "span B' = UNIV" - by (auto simp: top_unique) - have "\y. \X. {x. X x \ 0} \ B' \ finite {x. X x \ 0} \ y = (\x|X x \ 0. X x *\<^sub>R x)" - using \span B' = UNIV\ unfolding span_alt by auto - then obtain X where X: "\y. {x. X y x \ 0} \ B'" "\y. finite {x. X y x \ 0}" - "\y. y = (\x|X y x \ 0. X y x *\<^sub>R x)" - unfolding choice_iff by auto - - have X_add: "X (x + y) = (\z. X x z + X y z)" for x y - using \independent B'\ - proof (rule independentD_unique) - have "(\z | X x z + X y z \ 0. (X x z + X y z) *\<^sub>R z) - = (\z\{z. X x z \ 0} \ {z. X y z \ 0}. (X x z + X y z) *\<^sub>R z)" - by (intro sum.mono_neutral_cong_left) (auto intro: X) - also have "\ = (\z\{z. X x z \ 0}. X x z *\<^sub>R z) + (\z\{z. X y z \ 0}. X y z *\<^sub>R z)" - by (auto simp add: scaleR_add_left sum.distrib - intro!: arg_cong2[where f="(+)"] sum.mono_neutral_cong_right X) - also have "\ = x + y" - by (simp add: X(3)[symmetric]) - also have "\ = (\z | X (x + y) z \ 0. X (x + y) z *\<^sub>R z)" - by (rule X(3)) - finally show "(\z | X (x + y) z \ 0. X (x + y) z *\<^sub>R z) = (\z | X x z + X y z \ 0. (X x z + X y z) *\<^sub>R z)" - .. - have "{z. X x z + X y z \ 0} \ {z. X x z \ 0} \ {z. X y z \ 0}" - by auto - then show "finite {z. X x z + X y z \ 0}" "{xa. X x xa + X y xa \ 0} \ B'" - "finite {xa. X (x + y) xa \ 0}" "{xa. X (x + y) xa \ 0} \ B'" - using X(1) by (auto dest: finite_subset intro: X) - qed - - have X_cmult: "X (c *\<^sub>R x) = (\z. c * X x z)" for x c - using \independent B'\ - proof (rule independentD_unique) - show "finite {z. X (c *\<^sub>R x) z \ 0}" "{z. X (c *\<^sub>R x) z \ 0} \ B'" - "finite {z. c * X x z \ 0}" "{z. c * X x z \ 0} \ B' " - using X(1,2) by auto - show "(\z | X (c *\<^sub>R x) z \ 0. X (c *\<^sub>R x) z *\<^sub>R z) = (\z | c * X x z \ 0. (c * X x z) *\<^sub>R z)" - unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] - by (cases "c = 0") (auto simp: X(3)[symmetric]) - qed - - have X_B': "x \ B' \ X x = (\z. if z = x then 1 else 0)" for x - using \independent B'\ - by (rule independentD_unique[OF _ X(2) X(1)]) (auto intro: X simp: X(3)[symmetric]) - - define f' where "f' y = (if y \ B then f y else 0)" for y - define g where "g y = (\x|X y x \ 0. X y x *\<^sub>R f' x)" for y - - have g_f': "x \ B' \ g x = f' x" for x - by (auto simp: g_def X_B') - - have "linear g" - proof - fix x y - have *: "(\z | X x z + X y z \ 0. (X x z + X y z) *\<^sub>R f' z) - = (\z\{z. X x z \ 0} \ {z. X y z \ 0}. (X x z + X y z) *\<^sub>R f' z)" - by (intro sum.mono_neutral_cong_left) (auto intro: X) - show "g (x + y) = g x + g y" - unfolding g_def X_add * - by (auto simp add: scaleR_add_left sum.distrib - intro!: arg_cong2[where f="(+)"] sum.mono_neutral_cong_right X) - next - show "g (r *\<^sub>R x) = r *\<^sub>R g x" for r x - by (auto simp add: g_def X_cmult scaleR_sum_right intro!: sum.mono_neutral_cong_left X) - qed - moreover have "\x\B. g x = f x" - using \B \ B'\ by (auto simp: g_f' f'_def) - moreover have "range g = span (f`B)" - unfolding \span B' = UNIV\[symmetric] span_linear_image[OF \linear g\, symmetric] - proof (rule span_subspace) - have "g ` B' \ f`B \ {0}" - by (auto simp: g_f' f'_def) - also have "\ \ span (f`B)" - by (auto intro: span_superset span_0) - finally show "g ` B' \ span (f`B)" - by auto - have "x \ B \ f x = g x" for x - using \B \ B'\ by (auto simp add: g_f' f'_def) - then show "span (f ` B) \ span (g ` B')" - using \B \ B'\ by (intro span_mono) auto - qed (rule subspace_span) - ultimately show ?thesis - by auto -qed - -lemma linear_independent_extend: - "independent B \ \g. linear g \ (\x\B. g x = f x)" - using linear_independent_extend_subspace[of B f] by auto - -text \Linear functions are equal on a subspace if they are on a spanning set.\ - -lemma subspace_kernel: - assumes lf: "linear f" - shows "subspace {x. f x = 0}" - apply (simp add: subspace_def) - apply (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf]) - done - -lemma linear_eq_0_span: - assumes lf: "linear f" and f0: "\x\B. f x = 0" - shows "\x \ span B. f x = 0" - using f0 subspace_kernel[OF lf] - by (rule span_induct') - -lemma linear_eq_0: "linear f \ S \ span B \ \x\B. f x = 0 \ \x\S. f x = 0" - using linear_eq_0_span[of f B] by auto - -lemma linear_eq_span: "linear f \ linear g \ \x\B. f x = g x \ \x \ span B. f x = g x" - using linear_eq_0_span[of "\x. f x - g x" B] by (auto simp: linear_compose_sub) - -lemma linear_eq: "linear f \ linear g \ S \ span B \ \x\B. f x = g x \ \x\S. f x = g x" - using linear_eq_span[of f g B] by auto - -text \The degenerate case of the Exchange Lemma.\ - -lemma spanning_subset_independent: - assumes BA: "B \ A" - and iA: "independent A" - and AsB: "A \ span B" - shows "A = B" -proof - show "B \ A" by (rule BA) - - from span_mono[OF BA] span_mono[OF AsB] - have sAB: "span A = span B" unfolding span_span by blast - - { - fix x - assume x: "x \ A" - from iA have th0: "x \ span (A - {x})" - unfolding dependent_def using x by blast - from x have xsA: "x \ span A" - by (blast intro: span_superset) - have "A - {x} \ A" by blast - then have th1: "span (A - {x}) \ span A" - by (metis span_mono) - { - assume xB: "x \ B" - from xB BA have "B \ A - {x}" - by blast - then have "span B \ span (A - {x})" - by (metis span_mono) - with th1 th0 sAB have "x \ span A" - by blast - with x have False - by (metis span_superset) - } - then have "x \ B" by blast - } - then show "A \ B" by blast -qed - -text \Relation between bases and injectivity/surjectivity of map.\ - -lemma spanning_surjective_image: - assumes us: "UNIV \ span S" - and lf: "linear f" - and sf: "surj f" - shows "UNIV \ span (f ` S)" -proof - - have "UNIV \ f ` UNIV" - using sf by (auto simp add: surj_def) - also have " \ \ span (f ` S)" - using spans_image[OF lf us] . - finally show ?thesis . -qed - -lemma independent_inj_on_image: - assumes iS: "independent S" - and lf: "linear f" - and fi: "inj_on f (span S)" - shows "independent (f ` S)" -proof - - { - fix a - assume a: "a \ S" "f a \ span (f ` S - {f a})" - have eq: "f ` S - {f a} = f ` (S - {a})" - using fi \a\S\ by (auto simp add: inj_on_def span_superset) - from a have "f a \ f ` span (S - {a})" - unfolding eq span_linear_image[OF lf, of "S - {a}"] by blast - then have "a \ span (S - {a})" - by (rule inj_on_image_mem_iff_alt[OF fi, rotated]) - (insert span_mono[of "S - {a}" S], auto intro: span_superset \a\S\) - with a(1) iS have False - by (simp add: dependent_def) - } - then show ?thesis - unfolding dependent_def by blast -qed - -lemma independent_injective_image: - "independent S \ linear f \ inj f \ independent (f ` S)" - using independent_inj_on_image[of S f] by (auto simp: subset_inj_on) - -text \Detailed theorems about left and right invertibility in general case.\ - -lemma linear_inj_on_left_inverse: - assumes lf: "linear f" and fi: "inj_on f (span S)" - shows "\g. range g \ span S \ linear g \ (\x\span S. g (f x) = x)" -proof - - obtain B where "independent B" "B \ S" "S \ span B" - using maximal_independent_subset[of S] by auto - then have "span S = span B" - unfolding span_eq by (auto simp: span_superset) - with linear_independent_extend_subspace[OF independent_inj_on_image, OF \independent B\ lf] fi - obtain g where g: "linear g" "\x\f ` B. g x = inv_into B f x" "range g = span (inv_into B f ` f ` B)" - by fastforce - have fB: "inj_on f B" - using fi by (auto simp: \span S = span B\ intro: subset_inj_on span_superset) - - have "\x\span B. g (f x) = x" - proof (intro linear_eq_span) - show "linear (\x. x)" "linear (\x. g (f x))" - using linear_id linear_compose[OF \linear f\ \linear g\] by (auto simp: id_def comp_def) - show "\x \ B. g (f x) = x" - using g fi \span S = span B\ by (auto simp: fB) - qed - moreover - have "inv_into B f ` f ` B \ B" - by (auto simp: fB) - then have "range g \ span S" - unfolding g \span S = span B\ by (intro span_mono) - ultimately show ?thesis - using \span S = span B\ \linear g\ by auto -qed - -lemma linear_injective_left_inverse: "linear f \ inj f \ \g. linear g \ g \ f = id" - using linear_inj_on_left_inverse[of f UNIV] by (auto simp: fun_eq_iff span_UNIV) - -lemma linear_surj_right_inverse: - assumes lf: "linear f" and sf: "span T \ f`span S" - shows "\g. range g \ span S \ linear g \ (\x\span T. f (g x) = x)" -proof - - obtain B where "independent B" "B \ T" "T \ span B" - using maximal_independent_subset[of T] by auto - then have "span T = span B" - unfolding span_eq by (auto simp: span_superset) - - from linear_independent_extend_subspace[OF \independent B\, of "inv_into (span S) f"] - obtain g where g: "linear g" "\x\B. g x = inv_into (span S) f x" "range g = span (inv_into (span S) f`B)" - by auto - moreover have "x \ B \ f (inv_into (span S) f x) = x" for x - using \B \ T\ \span T \ f`span S\ by (intro f_inv_into_f) (auto intro: span_superset) - ultimately have "\x\B. f (g x) = x" - by auto - then have "\x\span B. f (g x) = x" - using linear_id linear_compose[OF \linear g\ \linear f\] - by (intro linear_eq_span) (auto simp: id_def comp_def) - moreover have "inv_into (span S) f ` B \ span S" - using \B \ T\ \span T \ f`span S\ by (auto intro: inv_into_into span_superset) - then have "range g \ span S" - unfolding g by (intro span_minimal subspace_span) auto - ultimately show ?thesis - using \linear g\ \span T = span B\ by auto -qed - -lemma linear_surjective_right_inverse: "linear f \ surj f \ \g. linear g \ f \ g = id" - using linear_surj_right_inverse[of f UNIV UNIV] - by (auto simp: span_UNIV fun_eq_iff) - -text \The general case of the Exchange Lemma, the key to what follows.\ - -lemma exchange_lemma: - assumes f:"finite t" - and i: "independent s" - and sp: "s \ span t" - shows "\t'. card t' = card t \ finite t' \ s \ t' \ t' \ s \ t \ s \ span t'" - using f i sp -proof (induct "card (t - s)" arbitrary: s t rule: less_induct) - case less - note ft = \finite t\ and s = \independent s\ and sp = \s \ span t\ - let ?P = "\t'. card t' = card t \ finite t' \ s \ t' \ t' \ s \ t \ s \ span t'" - let ?ths = "\t'. ?P t'" - { - assume "s \ t" - then have ?ths - by (metis ft Un_commute sp sup_ge1) - } - moreover - { - assume st: "t \ s" - from spanning_subset_independent[OF st s sp] st ft span_mono[OF st] - have ?ths - by (metis Un_absorb sp) - } - moreover - { - assume st: "\ s \ t" "\ t \ s" - from st(2) obtain b where b: "b \ t" "b \ s" - by blast - from b have "t - {b} - s \ t - s" - by blast - then have cardlt: "card (t - {b} - s) < card (t - s)" - using ft by (auto intro: psubset_card_mono) - from b ft have ct0: "card t \ 0" - by auto - have ?ths - proof cases - assume stb: "s \ span (t - {b})" - from ft have ftb: "finite (t - {b})" - by auto - from less(1)[OF cardlt ftb s stb] - obtain u where u: "card u = card (t - {b})" "s \ u" "u \ s \ (t - {b})" "s \ span u" - and fu: "finite u" by blast - let ?w = "insert b u" - have th0: "s \ insert b u" - using u by blast - from u(3) b have "u \ s \ t" - by blast - then have th1: "insert b u \ s \ t" - using u b by blast - have bu: "b \ u" - using b u by blast - from u(1) ft b have "card u = (card t - 1)" - by auto - then have th2: "card (insert b u) = card t" - using card_insert_disjoint[OF fu bu] ct0 by auto - from u(4) have "s \ span u" . - also have "\ \ span (insert b u)" - by (rule span_mono) blast - finally have th3: "s \ span (insert b u)" . - from th0 th1 th2 th3 fu have th: "?P ?w" - by blast - from th show ?thesis by blast - next - assume stb: "\ s \ span (t - {b})" - from stb obtain a where a: "a \ s" "a \ span (t - {b})" - by blast - have ab: "a \ b" - using a b by blast - have at: "a \ t" - using a ab span_superset[of a "t- {b}"] by auto - have mlt: "card ((insert a (t - {b})) - s) < card (t - s)" - using cardlt ft a b by auto - have ft': "finite (insert a (t - {b}))" - using ft by auto - { - fix x - assume xs: "x \ s" - have t: "t \ insert b (insert a (t - {b}))" - using b by auto - from b(1) have "b \ span t" - by (simp add: span_superset) - have bs: "b \ span (insert a (t - {b}))" - apply (rule in_span_delete) - using a sp unfolding subset_eq - apply auto - done - from xs sp have "x \ span t" - by blast - with span_mono[OF t] have x: "x \ span (insert b (insert a (t - {b})))" .. - from span_trans[OF bs x] have "x \ span (insert a (t - {b}))" . - } - then have sp': "s \ span (insert a (t - {b}))" - by blast - from less(1)[OF mlt ft' s sp'] obtain u where u: - "card u = card (insert a (t - {b}))" - "finite u" "s \ u" "u \ s \ insert a (t - {b})" - "s \ span u" by blast - from u a b ft at ct0 have "?P u" - by auto - then show ?thesis by blast - qed - } - ultimately show ?ths by blast -qed - -text \This implies corresponding size bounds.\ - -lemma independent_span_bound: - assumes f: "finite t" - and i: "independent s" - and sp: "s \ span t" - shows "finite s \ card s \ card t" - by (metis exchange_lemma[OF f i sp] finite_subset card_mono) - lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\ (UNIV::'a::finite set)}" proof - have eq: "{f x |x. x\ UNIV} = f ` UNIV" @@ -1353,51 +47,8 @@ subsection%unimportant \More interesting properties of the norm.\ -lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)" - by auto - notation inner (infix "\" 70) -lemma square_bound_lemma: - fixes x :: real - shows "x < (1 + x) * (1 + x)" -proof - - have "(x + 1/2)\<^sup>2 + 3/4 > 0" - using zero_le_power2[of "x+1/2"] by arith - then show ?thesis - by (simp add: field_simps power2_eq_square) -qed - -lemma square_continuous: - fixes e :: real - shows "e > 0 \ \d. 0 < d \ (\y. \y - x\ < d \ \y * y - x * x\ < e)" - using isCont_power[OF continuous_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2] - by (force simp add: power2_eq_square) - - -lemma norm_eq_0_dot: "norm x = 0 \ x \ x = (0::real)" - by simp (* TODO: delete *) - -lemma norm_triangle_sub: - fixes x y :: "'a::real_normed_vector" - shows "norm x \ norm y + norm (x - y)" - using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps) - -lemma norm_le: "norm x \ norm y \ x \ x \ y \ y" - by (simp add: norm_eq_sqrt_inner) - -lemma norm_lt: "norm x < norm y \ x \ x < y \ y" - by (simp add: norm_eq_sqrt_inner) - -lemma norm_eq: "norm x = norm y \ x \ x = y \ y" - apply (subst order_eq_iff) - apply (auto simp: norm_le) - done - -lemma norm_eq_1: "norm x = 1 \ x \ x = 1" - by (simp add: norm_eq_sqrt_inner) - - text\Equality of vectors in terms of @{term "(\)"} products.\ lemma linear_componentwise: @@ -1405,11 +56,11 @@ assumes lf: "linear f" shows "(f x) \ j = (\i\Basis. (x\i) * (f i\j))" (is "?lhs = ?rhs") proof - + interpret linear f by fact have "?rhs = (\i\Basis. (x\i) *\<^sub>R (f i))\j" by (simp add: inner_sum_left) then show ?thesis - unfolding linear_sum_mul[OF lf, symmetric] - unfolding euclidean_representation .. + by (simp add: euclidean_representation sum[symmetric] scale[symmetric]) qed lemma vector_eq: "x = y \ x \ x = x \ y \ y \ y = x \ x" @@ -1607,22 +258,15 @@ using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg) lemma bilinear_sum: - assumes bh: "bilinear h" - and fS: "finite S" - and fT: "finite T" + assumes "bilinear h" shows "h (sum f S) (sum g T) = sum (\(i,j). h (f i) (g j)) (S \ T) " proof - + interpret l: linear "\x. h x y" for y using assms by (simp add: bilinear_def) + interpret r: linear "\y. h x y" for x using assms by (simp add: bilinear_def) have "h (sum f S) (sum g T) = sum (\x. h (f x) (sum g T)) S" - apply (rule linear_sum[unfolded o_def]) - using bh fS - apply (auto simp add: bilinear_def) - done + by (simp add: l.sum) also have "\ = sum (\x. sum (\y. h (f x) (g y)) T) S" - apply (rule sum.cong, simp) - apply (rule linear_sum[unfolded o_def]) - using bh fT - apply (auto simp add: bilinear_def) - done + by (rule sum.cong) (simp_all add: r.sum) finally show ?thesis unfolding sum.cartesian_product . qed @@ -1663,6 +307,7 @@ assumes lf: "linear f" shows "x \ adjoint f y = f x \ y" proof - + interpret linear f by fact have "\y. \w. \x. f x \ y = x \ w" proof (intro allI exI) fix y :: "'m" and x @@ -1670,8 +315,7 @@ have "f x \ y = f (\i\Basis. (x \ i) *\<^sub>R i) \ y" by (simp add: euclidean_representation) also have "\ = (\i\Basis. (x \ i) *\<^sub>R f i) \ y" - unfolding linear_sum[OF lf] - by (simp add: linear_cmul[OF lf]) + by (simp add: sum scale) finally show "f x \ y = x \ ?w" by (simp add: inner_sum_left inner_sum_right mult.commute) qed @@ -1847,63 +491,14 @@ subsection%unimportant \Euclidean Spaces as Typeclass\ lemma independent_Basis: "independent Basis" - unfolding dependent_def - apply (subst span_finite) - apply simp - apply clarify - apply (drule_tac f="inner a" in arg_cong) - apply (simp add: inner_Basis inner_sum_right eq_commute) - done + by (rule independent_Basis) lemma span_Basis [simp]: "span Basis = UNIV" - unfolding span_finite [OF finite_Basis] - by (fast intro: euclidean_representation) + by (rule span_Basis) lemma in_span_Basis: "x \ span Basis" unfolding span_Basis .. -lemma Basis_le_norm: "b \ Basis \ \x \ b\ \ norm x" - by (rule order_trans [OF Cauchy_Schwarz_ineq2]) simp - -lemma norm_bound_Basis_le: "b \ Basis \ norm x \ e \ \x \ b\ \ e" - by (metis Basis_le_norm order_trans) - -lemma norm_bound_Basis_lt: "b \ Basis \ norm x < e \ \x \ b\ < e" - by (metis Basis_le_norm le_less_trans) - -lemma norm_le_l1: "norm x \ (\b\Basis. \x \ b\)" - apply (subst euclidean_representation[of x, symmetric]) - apply (rule order_trans[OF norm_sum]) - apply (auto intro!: sum_mono) - done - -lemma sum_norm_allsubsets_bound: - fixes f :: "'a \ 'n::euclidean_space" - assumes fP: "finite P" - and fPs: "\Q. Q \ P \ norm (sum f Q) \ e" - shows "(\x\P. norm (f x)) \ 2 * real DIM('n) * e" -proof - - have "(\x\P. norm (f x)) \ (\x\P. \b\Basis. \f x \ b\)" - by (rule sum_mono) (rule norm_le_l1) - also have "(\x\P. \b\Basis. \f x \ b\) = (\b\Basis. \x\P. \f x \ b\)" - by (rule sum.swap) - also have "\ \ of_nat (card (Basis :: 'n set)) * (2 * e)" - proof (rule sum_bounded_above) - fix i :: 'n - assume i: "i \ Basis" - have "norm (\x\P. \f x \ i\) \ - norm ((\x\P \ - {x. f x \ i < 0}. f x) \ i) + norm ((\x\P \ {x. f x \ i < 0}. f x) \ i)" - by (simp add: abs_real_def sum.If_cases[OF fP] sum_negf norm_triangle_ineq4 inner_sum_left - del: real_norm_def) - also have "\ \ e + e" - unfolding real_norm_def - by (intro add_mono norm_bound_Basis_le i fPs) auto - finally show "(\x\P. \f x \ i\) \ 2*e" by simp - qed - also have "\ = 2 * real DIM('n) * e" by simp - finally show ?thesis . -qed - subsection%unimportant \Linearity and Bilinearity continued\ @@ -1912,6 +507,7 @@ assumes lf: "linear f" shows "\B. \x. norm (f x) \ B * norm x" proof + interpret linear f by fact let ?B = "\b\Basis. norm (f b)" show "\x. norm (f x) \ ?B * norm x" proof @@ -1920,7 +516,7 @@ have "norm (f x) = norm (f (\b\Basis. (x \ b) *\<^sub>R b))" unfolding euclidean_representation .. also have "\ = norm (sum ?g Basis)" - by (simp add: linear_sum [OF lf] linear_cmul [OF lf]) + by (simp add: sum scale) finally have th0: "norm (f x) = norm (sum ?g Basis)" . have th: "norm (?g i) \ norm (f i) * norm x" if "i \ Basis" for i proof - @@ -1997,15 +593,15 @@ fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "linear f" "inj f" obtains B where "B > 0" "\x. B * norm x \ norm(f x)" - using linear_injective_left_inverse [OF assms] linear_invertible_bounded_below_pos assms by blast + using linear_injective_left_inverse [OF assms] + linear_invertible_bounded_below_pos assms by blast lemma bounded_linearI': fixes f ::"'a::euclidean_space \ 'b::real_normed_vector" assumes "\x y. f (x + y) = f x + f y" and "\c x. f (c *\<^sub>R x) = c *\<^sub>R f x" shows "bounded_linear f" - unfolding linear_conv_bounded_linear[symmetric] - by (rule linearI[OF assms]) + using assms linearI linear_conv_bounded_linear by blast lemma bilinear_bounded: fixes h :: "'m::euclidean_space \ 'n::euclidean_space \ 'k::real_normed_vector" @@ -2020,7 +616,7 @@ apply rule done also have "\ = norm (sum (\ (i,j). h ((x \ i) *\<^sub>R i) ((y \ j) *\<^sub>R j)) (Basis \ Basis))" - unfolding bilinear_sum[OF bh finite_Basis finite_Basis] .. + unfolding bilinear_sum[OF bh] .. finally have th: "norm (h x y) = \" . show "norm (h x y) \ (\i\Basis. \j\Basis. norm (h i j)) * norm x * norm y" apply (auto simp add: sum_distrib_right th sum.cartesian_product) @@ -2084,14 +680,14 @@ by (simp only: ac_simps) qed -lemma bounded_linear_imp_has_derivative: - "bounded_linear f \ (f has_derivative f) net" - by (simp add: has_derivative_def bounded_linear.linear linear_diff) +lemma bounded_linear_imp_has_derivative: "bounded_linear f \ (f has_derivative f) net" + by (auto simp add: has_derivative_def linear_diff linear_linear linear_def + dest: bounded_linear.linear) lemma linear_imp_has_derivative: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ (f has_derivative f) net" -by (simp add: has_derivative_def linear_conv_bounded_linear linear_diff) + by (simp add: bounded_linear_imp_has_derivative linear_conv_bounded_linear) lemma bounded_linear_imp_differentiable: "bounded_linear f \ f differentiable net" using bounded_linear_imp_has_derivative differentiable_def by blast @@ -2099,7 +695,7 @@ lemma linear_imp_differentiable: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable net" -by (metis linear_imp_has_derivative differentiable_def) + by (metis linear_imp_has_derivative differentiable_def) subsection%unimportant \We continue.\ @@ -2107,221 +703,21 @@ lemma independent_bound: fixes S :: "'a::euclidean_space set" shows "independent S \ finite S \ card S \ DIM('a)" - using independent_span_bound[OF finite_Basis, of S] by auto + by (metis dim_subset_UNIV finiteI_independent dim_span_eq_card_independent) + +lemmas independent_imp_finite = finiteI_independent corollary fixes S :: "'a::euclidean_space set" assumes "independent S" - shows independent_imp_finite: "finite S" and independent_card_le:"card S \ DIM('a)" -using assms independent_bound by auto - -lemma independent_explicit: - fixes B :: "'a::euclidean_space set" - shows "independent B \ - finite B \ (\c. (\v\B. c v *\<^sub>R v) = 0 \ (\v \ B. c v = 0))" -apply (cases "finite B") - apply (force simp: dependent_finite) -using independent_bound -apply auto -done + shows independent_card_le:"card S \ DIM('a)" + using assms independent_bound by auto lemma dependent_biggerset: fixes S :: "'a::euclidean_space set" shows "(finite S \ card S > DIM('a)) \ dependent S" by (metis independent_bound not_less) -text \Notion of dimension.\ - -definition "dim V = (SOME n. \B. B \ V \ independent B \ V \ span B \ card B = n)" - -lemma basis_exists: - "\B. (B :: ('a::euclidean_space) set) \ V \ independent B \ V \ span B \ (card B = dim V)" - unfolding dim_def some_eq_ex[of "\n. \B. B \ V \ independent B \ V \ span B \ (card B = n)"] - using maximal_independent_subset[of V] independent_bound - by auto - -corollary dim_le_card: - fixes s :: "'a::euclidean_space set" - shows "finite s \ dim s \ card s" -by (metis basis_exists card_mono) - -text \Consequences of independence or spanning for cardinality.\ - -lemma independent_card_le_dim: - fixes B :: "'a::euclidean_space set" - assumes "B \ V" - and "independent B" - shows "card B \ dim V" -proof - - from basis_exists[of V] \B \ V\ - obtain B' where "independent B'" - and "B \ span B'" - and "card B' = dim V" - by blast - with independent_span_bound[OF _ \independent B\ \B \ span B'\] independent_bound[of B'] - show ?thesis by auto -qed - -lemma span_card_ge_dim: - fixes B :: "'a::euclidean_space set" - shows "B \ V \ V \ span B \ finite B \ dim V \ card B" - by (metis basis_exists[of V] independent_span_bound subset_trans) - -lemma basis_card_eq_dim: - fixes V :: "'a::euclidean_space set" - shows "B \ V \ V \ span B \ independent B \ finite B \ card B = dim V" - by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_bound) - -lemma dim_unique: - fixes B :: "'a::euclidean_space set" - shows "B \ V \ V \ span B \ independent B \ card B = n \ dim V = n" - by (metis basis_card_eq_dim) - -text \More lemmas about dimension.\ - -lemma dim_UNIV [simp]: "dim (UNIV :: 'a::euclidean_space set) = DIM('a)" - using independent_Basis - by (intro dim_unique[of Basis]) auto - -lemma dim_subset: - fixes S :: "'a::euclidean_space set" - shows "S \ T \ dim S \ dim T" - using basis_exists[of T] basis_exists[of S] - by (metis independent_card_le_dim subset_trans) - -lemma dim_subset_UNIV: - fixes S :: "'a::euclidean_space set" - shows "dim S \ DIM('a)" - by (metis dim_subset subset_UNIV dim_UNIV) - -text \Converses to those.\ - -lemma card_ge_dim_independent: - fixes B :: "'a::euclidean_space set" - assumes BV: "B \ V" - and iB: "independent B" - and dVB: "dim V \ card B" - shows "V \ span B" -proof - fix a - assume aV: "a \ V" - { - assume aB: "a \ span B" - then have iaB: "independent (insert a B)" - using iB aV BV by (simp add: independent_insert) - from aV BV have th0: "insert a B \ V" - by blast - from aB have "a \B" - by (auto simp add: span_superset) - with independent_card_le_dim[OF th0 iaB] dVB independent_bound[OF iB] - have False by auto - } - then show "a \ span B" by blast -qed - -lemma card_le_dim_spanning: - assumes BV: "(B:: ('a::euclidean_space) set) \ V" - and VB: "V \ span B" - and fB: "finite B" - and dVB: "dim V \ card B" - shows "independent B" -proof - - { - fix a - assume a: "a \ B" "a \ span (B - {a})" - from a fB have c0: "card B \ 0" - by auto - from a fB have cb: "card (B - {a}) = card B - 1" - by auto - from BV a have th0: "B - {a} \ V" - by blast - { - fix x - assume x: "x \ V" - from a have eq: "insert a (B - {a}) = B" - by blast - from x VB have x': "x \ span B" - by blast - from span_trans[OF a(2), unfolded eq, OF x'] - have "x \ span (B - {a})" . - } - then have th1: "V \ span (B - {a})" - by blast - have th2: "finite (B - {a})" - using fB by auto - from span_card_ge_dim[OF th0 th1 th2] - have c: "dim V \ card (B - {a})" . - from c c0 dVB cb have False by simp - } - then show ?thesis - unfolding dependent_def by blast -qed - -lemma card_eq_dim: - fixes B :: "'a::euclidean_space set" - shows "B \ V \ card B = dim V \ finite B \ independent B \ V \ span B" - by (metis order_eq_iff card_le_dim_spanning card_ge_dim_independent) - -text \More general size bound lemmas.\ - -lemma independent_bound_general: - fixes S :: "'a::euclidean_space set" - shows "independent S \ finite S \ card S \ dim S" - by (metis independent_card_le_dim independent_bound subset_refl) - -lemma dependent_biggerset_general: - fixes S :: "'a::euclidean_space set" - shows "(finite S \ card S > dim S) \ dependent S" - using independent_bound_general[of S] by (metis linorder_not_le) - -lemma dim_span [simp]: - fixes S :: "'a::euclidean_space set" - shows "dim (span S) = dim S" -proof - - have th0: "dim S \ dim (span S)" - by (auto simp add: subset_eq intro: dim_subset span_superset) - from basis_exists[of S] - obtain B where B: "B \ S" "independent B" "S \ span B" "card B = dim S" - by blast - from B have fB: "finite B" "card B = dim S" - using independent_bound by blast+ - have bSS: "B \ span S" - using B(1) by (metis subset_eq span_inc) - have sssB: "span S \ span B" - using span_mono[OF B(3)] by (simp add: span_span) - from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis - using fB(2) by arith -qed - -lemma subset_le_dim: - fixes S :: "'a::euclidean_space set" - shows "S \ span T \ dim S \ dim T" - by (metis dim_span dim_subset) - -lemma span_eq_dim: - fixes S :: "'a::euclidean_space set" - shows "span S = span T \ dim S = dim T" - by (metis dim_span) - -lemma dim_image_le: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes lf: "linear f" - shows "dim (f ` S) \ dim (S)" -proof - - from basis_exists[of S] obtain B where - B: "B \ S" "independent B" "S \ span B" "card B = dim S" by blast - from B have fB: "finite B" "card B = dim S" - using independent_bound by blast+ - have "dim (f ` S) \ card (f ` B)" - apply (rule span_card_ge_dim) - using lf B fB - apply (auto simp add: span_linear_image spans_image subset_image_iff) - done - also have "\ \ dim S" - using card_image_le[OF fB(1)] fB by simp - finally show ?thesis . -qed - text \Picking an orthogonal replacement for a spanning set.\ lemma vector_sub_project_orthogonal: @@ -2367,10 +763,10 @@ have "x - k *\<^sub>R (a - (\x\C. (x \ a / (x \ x)) *\<^sub>R x)) \ span C \ x - k *\<^sub>R a \ span C" apply (simp only: scaleR_right_diff_distrib th0) apply (rule span_add_eq) - apply (rule span_mul) + apply (rule span_scale) apply (rule span_sum) - apply (rule span_mul) - apply (rule span_superset) + apply (rule span_scale) + apply (rule span_base) apply assumption done } @@ -2402,7 +798,8 @@ lemma orthogonal_basis_exists: fixes V :: "('a::euclidean_space) set" - shows "\B. independent B \ B \ span V \ V \ span B \ (card B = dim V) \ pairwise orthogonal B" + shows "\B. independent B \ B \ span V \ V \ span B \ + (card B = dim V) \ pairwise orthogonal B" proof - from basis_exists[of V] obtain B where B: "B \ V" "independent B" "V \ span B" "card B = dim V" @@ -2413,7 +810,7 @@ C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast from C B have CSV: "C \ span V" - by (metis span_inc span_mono subset_trans) + by (metis span_superset span_mono subset_trans) from span_mono[OF B(3)] C have SVC: "span V \ span C" by (simp add: span_span) from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB @@ -2423,7 +820,7 @@ by simp moreover have "dim V \ card C" using span_card_ge_dim[OF CSV SVC C(1)] - by (simp add: dim_span) + by simp ultimately have CdV: "card C = dim V" using C(1) by simp from C B CSV CdV iC show ?thesis @@ -2440,7 +837,8 @@ from sU obtain a where a: "a \ span S" by blast from orthogonal_basis_exists obtain B where - B: "independent B" "B \ span S" "S \ span B" "card B = dim S" "pairwise orthogonal B" + B: "independent B" "B \ span S" "S \ span B" + "card B = dim S" "pairwise orthogonal B" by blast from B have fB: "finite B" "card B = dim S" using independent_bound by auto @@ -2451,8 +849,8 @@ have "sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B \ span S" unfolding sSB apply (rule span_sum) - apply (rule span_mul) - apply (rule span_superset) + apply (rule span_scale) + apply (rule span_base) apply assumption done with a have a0:"?a \ 0" @@ -2504,7 +902,7 @@ then have "dim (span S) = dim (UNIV :: ('a) set)" by simp then have "dim S = DIM('a)" - by (simp add: dim_span dim_UNIV) + by (metis Euclidean_Space.dim_UNIV dim_span) with d have False by arith } then have th: "span S \ UNIV" @@ -2512,132 +910,15 @@ from span_not_univ_subset_hyperplane[OF th] show ?thesis . qed -text \We can extend a linear basis-basis injection to the whole set.\ - -lemma linear_indep_image_lemma: - assumes lf: "linear f" - and fB: "finite B" - and ifB: "independent (f ` B)" - and fi: "inj_on f B" - and xsB: "x \ span B" - and fx: "f x = 0" - shows "x = 0" - using fB ifB fi xsB fx -proof (induct arbitrary: x rule: finite_induct[OF fB]) - case 1 - then show ?case by auto -next - case (2 a b x) - have fb: "finite b" using "2.prems" by simp - have th0: "f ` b \ f ` (insert a b)" - apply (rule image_mono) - apply blast - done - from independent_mono[ OF "2.prems"(2) th0] - have ifb: "independent (f ` b)" . - have fib: "inj_on f b" - apply (rule subset_inj_on [OF "2.prems"(3)]) - apply blast - done - from span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)] - obtain k where k: "x - k*\<^sub>R a \ span (b - {a})" - by blast - have "f (x - k*\<^sub>R a) \ span (f ` b)" - unfolding span_linear_image[OF lf] - apply (rule imageI) - using k span_mono[of "b - {a}" b] - apply blast - done - then have "f x - k*\<^sub>R f a \ span (f ` b)" - by (simp add: linear_diff[OF lf] linear_cmul[OF lf]) - then have th: "-k *\<^sub>R f a \ span (f ` b)" - using "2.prems"(5) by simp - have xsb: "x \ span b" - proof (cases "k = 0") - case True - with k have "x \ span (b - {a})" by simp - then show ?thesis using span_mono[of "b - {a}" b] - by blast - next - case False - with span_mul[OF th, of "- 1/ k"] - have th1: "f a \ span (f ` b)" - by auto - from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric] - have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast - from "2.prems"(2) [unfolded dependent_def bex_simps(8), rule_format, of "f a"] - have "f a \ span (f ` b)" using tha - using "2.hyps"(2) - "2.prems"(3) by auto - with th1 have False by blast - then show ?thesis by blast - qed - from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)] show "x = 0" . -qed - -text \Can construct an isomorphism between spaces of same dimension.\ - -lemma subspace_isomorphism: - fixes S :: "'a::euclidean_space set" - and T :: "'b::euclidean_space set" - assumes s: "subspace S" - and t: "subspace T" - and d: "dim S = dim T" - shows "\f. linear f \ f ` S = T \ inj_on f S" -proof - - from basis_exists[of S] independent_bound - obtain B where B: "B \ S" "independent B" "S \ span B" "card B = dim S" and fB: "finite B" - by blast - from basis_exists[of T] independent_bound - obtain C where C: "C \ T" "independent C" "T \ span C" "card C = dim T" and fC: "finite C" - by blast - from B(4) C(4) card_le_inj[of B C] d - obtain f where f: "f ` B \ C" "inj_on f B" using \finite B\ \finite C\ - by auto - from linear_independent_extend[OF B(2)] - obtain g where g: "linear g" "\x\ B. g x = f x" - by blast - from inj_on_iff_eq_card[OF fB, of f] f(2) have "card (f ` B) = card B" - by simp - with B(4) C(4) have ceq: "card (f ` B) = card C" - using d by simp - have "g ` B = f ` B" - using g(2) by (auto simp add: image_iff) - also have "\ = C" using card_subset_eq[OF fC f(1) ceq] . - finally have gBC: "g ` B = C" . - have gi: "inj_on g B" - using f(2) g(2) by (auto simp add: inj_on_def) - note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] - { - fix x y - assume x: "x \ S" and y: "y \ S" and gxy: "g x = g y" - from B(3) x y have x': "x \ span B" and y': "y \ span B" - by blast+ - from gxy have th0: "g (x - y) = 0" - by (simp add: linear_diff[OF g(1)]) - have th1: "x - y \ span B" - using x' y' by (metis span_diff) - have "x = y" - using g0[OF th1 th0] by simp - } - then have giS: "inj_on g S" - unfolding inj_on_def by blast - from span_subspace[OF B(1,3) s] have "g ` S = span (g ` B)" - by (simp add: span_linear_image[OF g(1)]) - also have "\ = span C" unfolding gBC .. - also have "\ = T" using span_subspace[OF C(1,3) t] . - finally have gS: "g ` S = T" . - from g(1) gS giS show ?thesis - by blast -qed - lemma linear_eq_stdbasis: fixes f :: "'a::euclidean_space \ _" assumes lf: "linear f" and lg: "linear g" - and fg: "\b\Basis. f b = g b" + and fg: "\b. b \ Basis \ f b = g b" shows "f = g" - using linear_eq[OF lf lg, of _ Basis] fg by auto + using linear_eq_on_span[OF lf lg, of Basis] fg + by auto + text \Similar results for bilinear functions.\ @@ -2652,7 +933,8 @@ let ?P = "{x. \y\ span C. f x y = g x y}" from bf bg have sp: "subspace ?P" unfolding bilinear_def linear_iff subspace_def bf bg - by (auto simp add: span_0 bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def + by (auto simp add: span_zero bilinear_lzero[OF bf] bilinear_lzero[OF bg] + span_add Ball_def intro: bilinear_ladd[OF bf]) have "\x \ span B. \y\ span C. f x y = g x y" @@ -2662,7 +944,8 @@ apply (simp add: fg) apply (auto simp add: subspace_def) using bf bg unfolding bilinear_def linear_iff - apply (auto simp add: span_0 bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def + apply (auto simp add: span_zero bilinear_rzero[OF bf] bilinear_rzero[OF bg] + span_add Ball_def intro: bilinear_ladd[OF bf]) done then show ?thesis @@ -2677,234 +960,6 @@ shows "f = g" using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis] fg] by blast -text \An injective map @{typ "'a::euclidean_space \ 'b::euclidean_space"} is also surjective.\ - -lemma linear_injective_imp_surjective: - fixes f :: "'a::euclidean_space \ 'a::euclidean_space" - assumes lf: "linear f" - and fi: "inj f" - shows "surj f" -proof - - let ?U = "UNIV :: 'a set" - from basis_exists[of ?U] obtain B - where B: "B \ ?U" "independent B" "?U \ span B" "card B = dim ?U" - by blast - from B(4) have d: "dim ?U = card B" - by simp - have th: "?U \ span (f ` B)" - apply (rule card_ge_dim_independent) - apply blast - apply (rule independent_injective_image[OF B(2) lf fi]) - apply (rule order_eq_refl) - apply (rule sym) - unfolding d - apply (rule card_image) - apply (rule subset_inj_on[OF fi]) - apply blast - done - from th show ?thesis - unfolding span_linear_image[OF lf] surj_def - using B(3) by blast -qed - -text \And vice versa.\ - -lemma surjective_iff_injective_gen: - assumes fS: "finite S" - and fT: "finite T" - and c: "card S = card T" - and ST: "f ` S \ T" - shows "(\y \ T. \x \ S. f x = y) \ inj_on f S" - (is "?lhs \ ?rhs") -proof - assume h: "?lhs" - { - fix x y - assume x: "x \ S" - assume y: "y \ S" - assume f: "f x = f y" - from x fS have S0: "card S \ 0" - by auto - have "x = y" - proof (rule ccontr) - assume xy: "\ ?thesis" - have th: "card S \ card (f ` (S - {y}))" - unfolding c - apply (rule card_mono) - apply (rule finite_imageI) - using fS apply simp - using h xy x y f unfolding subset_eq image_iff - apply auto - apply (case_tac "xa = f x") - apply (rule bexI[where x=x]) - apply auto - done - also have " \ \ card (S - {y})" - apply (rule card_image_le) - using fS by simp - also have "\ \ card S - 1" using y fS by simp - finally show False using S0 by arith - qed - } - then show ?rhs - unfolding inj_on_def by blast -next - assume h: ?rhs - have "f ` S = T" - apply (rule card_subset_eq[OF fT ST]) - unfolding card_image[OF h] - apply (rule c) - done - then show ?lhs by blast -qed - -lemma linear_surjective_imp_injective: - fixes f :: "'a::euclidean_space \ 'a::euclidean_space" - assumes lf: "linear f" - and sf: "surj f" - shows "inj f" -proof - - let ?U = "UNIV :: 'a set" - from basis_exists[of ?U] obtain B - where B: "B \ ?U" "independent B" "?U \ span B" and d: "card B = dim ?U" - by blast - { - fix x - assume x: "x \ span B" - assume fx: "f x = 0" - from B(2) have fB: "finite B" - using independent_bound by auto - have fBi: "independent (f ` B)" - apply (rule card_le_dim_spanning[of "f ` B" ?U]) - apply blast - using sf B(3) - unfolding span_linear_image[OF lf] surj_def subset_eq image_iff - apply blast - using fB apply blast - unfolding d[symmetric] - apply (rule card_image_le) - apply (rule fB) - done - have th0: "dim ?U \ card (f ` B)" - apply (rule span_card_ge_dim) - apply blast - unfolding span_linear_image[OF lf] - apply (rule subset_trans[where B = "f ` UNIV"]) - using sf unfolding surj_def - apply blast - apply (rule image_mono) - apply (rule B(3)) - apply (metis finite_imageI fB) - done - moreover have "card (f ` B) \ card B" - by (rule card_image_le, rule fB) - ultimately have th1: "card B = card (f ` B)" - unfolding d by arith - have fiB: "inj_on f B" - unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] - by blast - from linear_indep_image_lemma[OF lf fB fBi fiB x] fx - have "x = 0" by blast - } - then show ?thesis - unfolding linear_injective_0[OF lf] - using B(3) - by blast -qed - -text \Hence either is enough for isomorphism.\ - -lemma left_right_inverse_eq: - assumes fg: "f \ g = id" - and gh: "g \ h = id" - shows "f = h" -proof - - have "f = f \ (g \ h)" - unfolding gh by simp - also have "\ = (f \ g) \ h" - by (simp add: o_assoc) - finally show "f = h" - unfolding fg by simp -qed - -lemma isomorphism_expand: - "f \ g = id \ g \ f = id \ (\x. f (g x) = x) \ (\x. g (f x) = x)" - by (simp add: fun_eq_iff o_def id_def) - -lemma linear_injective_isomorphism: - fixes f :: "'a::euclidean_space \ 'a::euclidean_space" - assumes lf: "linear f" - and fi: "inj f" - shows "\f'. linear f' \ (\x. f' (f x) = x) \ (\x. f (f' x) = x)" - unfolding isomorphism_expand[symmetric] - using linear_surjective_right_inverse[OF lf linear_injective_imp_surjective[OF lf fi]] - linear_injective_left_inverse[OF lf fi] - by (metis left_right_inverse_eq) - -lemma linear_surjective_isomorphism: - fixes f :: "'a::euclidean_space \ 'a::euclidean_space" - assumes lf: "linear f" - and sf: "surj f" - shows "\f'. linear f' \ (\x. f' (f x) = x) \ (\x. f (f' x) = x)" - unfolding isomorphism_expand[symmetric] - using linear_surjective_right_inverse[OF lf sf] - linear_injective_left_inverse[OF lf linear_surjective_imp_injective[OF lf sf]] - by (metis left_right_inverse_eq) - -text \Left and right inverses are the same for - @{typ "'a::euclidean_space \ 'a::euclidean_space"}.\ - -lemma linear_inverse_left: - fixes f :: "'a::euclidean_space \ 'a::euclidean_space" - assumes lf: "linear f" - and lf': "linear f'" - shows "f \ f' = id \ f' \ f = id" -proof - - { - fix f f':: "'a \ 'a" - assume lf: "linear f" "linear f'" - assume f: "f \ f' = id" - from f have sf: "surj f" - apply (auto simp add: o_def id_def surj_def) - apply metis - done - from linear_surjective_isomorphism[OF lf(1) sf] lf f - have "f' \ f = id" - unfolding fun_eq_iff o_def id_def by metis - } - then show ?thesis - using lf lf' by metis -qed - -text \Moreover, a one-sided inverse is automatically linear.\ - -lemma left_inverse_linear: - fixes f :: "'a::euclidean_space \ 'a::euclidean_space" - assumes lf: "linear f" - and gf: "g \ f = id" - shows "linear g" -proof - - from gf have fi: "inj f" - apply (auto simp add: inj_on_def o_def id_def fun_eq_iff) - apply metis - done - from linear_injective_isomorphism[OF lf fi] - obtain h :: "'a \ 'a" where h: "linear h" "\x. h (f x) = x" "\x. f (h x) = x" - by blast - have "h = g" - apply (rule ext) using gf h(2,3) - apply (simp add: o_def id_def fun_eq_iff) - apply metis - done - with h(1) show ?thesis by blast -qed - -lemma inj_linear_imp_inv_linear: - fixes f :: "'a::euclidean_space \ 'a::euclidean_space" - assumes "linear f" "inj f" shows "linear (inv f)" -using assms inj_iff left_inverse_linear by blast - - subsection \Infinity norm\ definition%important "infnorm (x::'a::euclidean_space) = Sup {\x \ b\ |b. b \ Basis}" @@ -3181,7 +1236,7 @@ qed then show ?thesis apply (clarsimp simp: collinear_def) - by (metis real_vector.scale_zero_right vector_fraction_eq_iff) + by (metis scaleR_zero_right vector_fraction_eq_iff) qed lemma collinear_subset: "\collinear T; S \ T\ \ collinear S" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Path_Connected.thy Wed May 02 13:49:38 2018 +0200 @@ -6354,21 +6354,18 @@ by (meson subsetD image_eqI inj_on_def) obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" using linear_independent_extend \independent B\ by fastforce - have "f ` S \ T" - by (metis ffb \fb ` B \ C\ \linear f\ \span B = S\ \span C = T\ image_cong span_linear_image span_mono) + have "span (f ` B) \ span C" + by (metis \fb ` B \ C\ ffb image_cong span_mono) + then have "f ` S \ T" + unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . have [simp]: "\x. x \ B \ norm (fb x) = norm x" using B1 C1 \fb ` B \ C\ by auto have "norm (f x) = norm x" if "x \ S" for x proof - + interpret linear f by fact obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce - have "f x = (\v \ B. f (a v *\<^sub>R v))" - using linear_sum [OF \linear f\] x by auto - also have "... = (\v \ B. a v *\<^sub>R f v)" - using \linear f\ by (simp add: linear_sum linear.scaleR) - also have "... = (\v \ B. a v *\<^sub>R fb v)" - by (simp add: ffb cong: sum.cong) - finally have "norm (f x)^2 = norm (\v\B. a v *\<^sub>R fb v)^2" by simp + have "norm (f x)^2 = norm (\v\B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) also have "... = (\v\B. norm ((a v *\<^sub>R fb v))^2)" apply (rule norm_sum_Pythagorean [OF \finite B\]) apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) @@ -6410,6 +6407,7 @@ by (meson subsetD image_eqI inj_on_def) obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" using linear_independent_extend \independent B\ by fastforce + interpret f: linear f by fact define gb where "gb \ inv_into B fb" then have pairwise_orth_gb: "pairwise (\v j. orthogonal (gb v) (gb j)) C" using Borth @@ -6417,8 +6415,12 @@ by (metis \bij_betw fb B C\ bij_betw_imp_surj_on bij_betw_inv_into_right inv_into_into) obtain g where "linear g" and ggb: "\x. x \ C \ g x = gb x" using linear_independent_extend \independent C\ by fastforce - have "f ` S \ T" - by (metis \bij_betw fb B C\ bij_betw_imp_surj_on eq_iff ffb \linear f\ \span B = S\ \span C = T\ image_cong span_linear_image) + interpret g: linear g by fact + have "span (f ` B) \ span C" + by (metis \bij_betw fb B C\ bij_betw_imp_surj_on eq_iff ffb image_cong) + then have "f ` S \ T" + unfolding \span B = S\ \span C = T\ + span_linear_image[OF \linear f\] . have [simp]: "\x. x \ B \ norm (fb x) = norm x" using B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on by fastforce have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \ S" for x @@ -6428,7 +6430,7 @@ have "f x = (\v \ B. f (a v *\<^sub>R v))" using linear_sum [OF \linear f\] x by auto also have "... = (\v \ B. a v *\<^sub>R f v)" - using \linear f\ by (simp add: linear_sum linear.scaleR) + by (simp add: f.sum f.scale) also have "... = (\v \ B. a v *\<^sub>R fb v)" by (simp add: ffb cong: sum.cong) finally have *: "f x = (\v\B. a v *\<^sub>R fb v)" . @@ -6443,9 +6445,9 @@ by (simp add: norm_eq_sqrt_inner) have "g (f x) = g (\v\B. a v *\<^sub>R fb v)" by (simp add: *) also have "... = (\v\B. g (a v *\<^sub>R fb v))" - using \linear g\ by (simp add: linear_sum linear.scaleR) + by (simp add: g.sum g.scale) also have "... = (\v\B. a v *\<^sub>R g (fb v))" - by (simp add: \linear g\ linear.scaleR) + by (simp add: g.scale) also have "... = (\v\B. a v *\<^sub>R v)" apply (rule sum.cong [OF refl]) using \bij_betw fb B C\ gb_def bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce @@ -6460,16 +6462,16 @@ obtain a where x: "x = (\v \ C. a v *\<^sub>R v)" using \finite C\ \span C = T\ \x \ T\ span_finite by fastforce have "g x = (\v \ C. g (a v *\<^sub>R v))" - using linear_sum [OF \linear g\] x by auto + by (simp add: x g.sum) also have "... = (\v \ C. a v *\<^sub>R g v)" - using \linear g\ by (simp add: linear_sum linear.scaleR) + by (simp add: g.scale) also have "... = (\v \ C. a v *\<^sub>R gb v)" by (simp add: ggb cong: sum.cong) finally have "f (g x) = f (\v\C. a v *\<^sub>R gb v)" by simp also have "... = (\v\C. f (a v *\<^sub>R gb v))" - using \linear f\ by (simp add: linear_sum linear.scaleR) + by (simp add: f.scale f.sum) also have "... = (\v\C. a v *\<^sub>R f (gb v))" - by (simp add: \linear f\ linear.scaleR) + by (simp add: f.scale f.sum) also have "... = (\v\C. a v *\<^sub>R v)" using \bij_betw fb B C\ by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) @@ -6478,7 +6480,8 @@ finally show "f (g x) = x" . qed have gim: "g ` T = S" - by (metis (no_types, lifting) \f ` S \ T\ \linear g\ \span B = S\ \span C = T\ d dim_eq_span dim_image_le f(2) image_subset_iff span_linear_image span_span subsetI) + by (metis (full_types) S T \f ` S \ T\ d dim_eq_span dim_image_le f(2) g.linear_axioms + image_iff linear_subspace_image span_eq_iff subset_iff) have fim: "f ` S = T" using \g ` T = S\ image_iff by fastforce have [simp]: "norm (g x) = norm x" if "x \ T" for x @@ -6504,8 +6507,8 @@ obtains f::"'M::euclidean_space \'N::euclidean_space" and g where "linear f" "linear g" "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" - "\x. g(f x) = x" "\y. f(g y) = y" - using assms by (auto simp: dim_UNIV intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) + "\x. g (f x) = x" "\y. f(g y) = y" + using assms by (auto simp: intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) lemma homeomorphic_subspaces: fixes S :: "'a::euclidean_space set" @@ -7079,7 +7082,7 @@ shows "uncountable (ball a r)" proof - have "uncountable (open_segment a (a + r *\<^sub>R (SOME i. i \ Basis)))" - by (metis Basis_zero SOME_Basis add_cancel_right_right assms less_le real_vector.scale_eq_0_iff uncountable_open_segment) + by (metis Basis_zero SOME_Basis add_cancel_right_right assms less_le scale_eq_0_iff uncountable_open_segment) moreover have "open_segment a (a + r *\<^sub>R (SOME i. i \ Basis)) \ ball a r" using assms by (auto simp: in_segment algebra_simps dist_norm SOME_Basis) ultimately show ?thesis diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Polytope.thy Wed May 02 13:49:38 2018 +0200 @@ -197,8 +197,7 @@ by (simp add: algebra_simps d_def) (simp add: divide_simps) have [simp]: "((e - e * e / (e + norm (b - c))) / norm (b - c)) = (e / (e + norm (b - c)))" using False nbc - apply (simp add: algebra_simps divide_simps) - by (metis mult_eq_0_iff norm_eq_zero norm_imp_pos_and_ge norm_pths(2) real_scaleR_def scaleR_left.add zero_less_norm_iff) + by (simp add: divide_simps) (simp add: algebra_simps) have "b \ open_segment d c" apply (simp add: open_segment_image_interval) apply (simp add: d_def algebra_simps image_def) @@ -2673,7 +2672,8 @@ lemma polyhedron_negations: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ polyhedron(image uminus S)" -by (auto simp: polyhedron_linear_image_eq linear_uminus bij_uminus) + by (subst polyhedron_linear_image_eq) + (auto simp: bij_uminus intro!: linear_uminus) subsection\Relation between polytopes and polyhedra\ diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Product_Vector.thy --- a/src/HOL/Analysis/Product_Vector.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Product_Vector.thy Wed May 02 13:49:38 2018 +0200 @@ -10,6 +10,52 @@ "HOL-Library.Product_Plus" begin +lemma Times_eq_image_sum: + fixes S :: "'a :: comm_monoid_add set" and T :: "'b :: comm_monoid_add set" + shows "S \ T = {u + v |u v. u \ (\x. (x, 0)) ` S \ v \ Pair 0 ` T}" + by force + + +subsection \Product is a module\ + +locale module_prod = module_pair begin + +definition scale :: "'a \ 'b \ 'c \ 'b \ 'c" + where "scale a v = (s1 a (fst v), s2 a (snd v))" + +lemma scale_prod: "scale x (a, b) = (s1 x a, s2 x b)" + by (auto simp: scale_def) + +sublocale p: module scale +proof qed (simp_all add: scale_def + m1.scale_left_distrib m1.scale_right_distrib m2.scale_left_distrib m2.scale_right_distrib) + +lemma subspace_Times: "m1.subspace A \ m2.subspace B \ p.subspace (A \ B)" + unfolding m1.subspace_def m2.subspace_def p.subspace_def + by (auto simp: zero_prod_def scale_def) + +lemma module_hom_fst: "module_hom scale s1 fst" + by unfold_locales (auto simp: scale_def) + +lemma module_hom_snd: "module_hom scale s2 snd" + by unfold_locales (auto simp: scale_def) + +end + +locale vector_space_prod = vector_space_pair begin + +sublocale module_prod s1 s2 + rewrites "module_hom = Vector_Spaces.linear" + by unfold_locales (fact module_hom_eq_linear) + +sublocale p: vector_space scale by unfold_locales (auto simp: algebra_simps) + +lemmas linear_fst = module_hom_fst + and linear_snd = module_hom_snd + +end + + subsection \Product is a real vector space\ instantiation%important prod :: (real_vector, real_vector) real_vector @@ -42,6 +88,27 @@ end +lemma module_prod_scale_eq_scaleR: "module_prod.scale ( *\<^sub>R) ( *\<^sub>R) = scaleR" + apply (rule ext) apply (rule ext) + apply (subst module_prod.scale_def) + subgoal by unfold_locales + by (simp add: scaleR_prod_def) + +interpretation real_vector?: vector_space_prod "scaleR::_\_\'a::real_vector" "scaleR::_\_\'b::real_vector" + rewrites "scale = (( *\<^sub>R)::_\_\('a \ 'b))" + and "module.dependent ( *\<^sub>R) = dependent" + and "module.representation ( *\<^sub>R) = representation" + and "module.subspace ( *\<^sub>R) = subspace" + and "module.span ( *\<^sub>R) = span" + and "vector_space.extend_basis ( *\<^sub>R) = extend_basis" + and "vector_space.dim ( *\<^sub>R) = dim" + and "Vector_Spaces.linear ( *\<^sub>R) ( *\<^sub>R) = linear" + subgoal by unfold_locales + subgoal by (fact module_prod_scale_eq_scaleR) + unfolding dependent_raw_def representation_raw_def subspace_raw_def span_raw_def + extend_basis_raw_def dim_raw_def linear_def + by (rule refl)+ + subsection \Product is a metric space\ (* TODO: Product of uniform spaces and compatibility with metric_spaces! *) @@ -270,7 +337,7 @@ show "(f (x + y), g (x + y)) = (f x, g x) + (f y, g y)" by (simp add: f.add g.add) show "(f (r *\<^sub>R x), g (r *\<^sub>R x)) = r *\<^sub>R (f x, g x)" - by (simp add: f.scaleR g.scaleR) + by (simp add: f.scale g.scale) obtain Kf where "0 < Kf" and norm_f: "\x. norm (f x) \ norm x * Kf" using f.pos_bounded by fast obtain Kg where "0 < Kg" and norm_g: "\x. norm (g x) \ norm x * Kg" @@ -389,4 +456,112 @@ unfolding norm_Pair by (metis norm_ge_zero sqrt_sum_squares_le_sum) +lemma (in vector_space_prod) span_Times_sing1: "p.span ({0} \ B) = {0} \ vs2.span B" + apply (rule p.span_unique) + subgoal by (auto intro!: vs1.span_base vs2.span_base) + subgoal using vs1.subspace_single_0 vs2.subspace_span by (rule subspace_Times) + subgoal for T + proof safe + fix b + assume subset_T: "{0} \ B \ T" and subspace: "p.subspace T" and b_span: "b \ vs2.span B" + then obtain t r where b: "b = (\a\t. r a *b a)" and t: "finite t" "t \ B" + by (auto simp: vs2.span_explicit) + have "(0, b) = (\b\t. scale (r b) (0, b))" + unfolding b scale_prod sum_prod + by simp + also have "\ \ T" + using \t \ B\ subset_T + by (auto intro!: p.subspace_sum p.subspace_scale subspace) + finally show "(0, b) \ T" . + qed + done + +lemma (in vector_space_prod) span_Times_sing2: "p.span (A \ {0}) = vs1.span A \ {0}" + apply (rule p.span_unique) + subgoal by (auto intro!: vs1.span_base vs2.span_base) + subgoal using vs1.subspace_span vs2.subspace_single_0 by (rule subspace_Times) + subgoal for T + proof safe + fix a + assume subset_T: "A \ {0} \ T" and subspace: "p.subspace T" and a_span: "a \ vs1.span A" + then obtain t r where a: "a = (\a\t. r a *a a)" and t: "finite t" "t \ A" + by (auto simp: vs1.span_explicit) + have "(a, 0) = (\a\t. scale (r a) (a, 0))" + unfolding a scale_prod sum_prod + by simp + also have "\ \ T" + using \t \ A\ subset_T + by (auto intro!: p.subspace_sum p.subspace_scale subspace) + finally show "(a, 0) \ T" . + qed + done + +lemma (in finite_dimensional_vector_space) zero_not_in_Basis[simp]: "0 \ Basis" + using dependent_zero local.independent_Basis by blast + +locale finite_dimensional_vector_space_prod = vector_space_prod + finite_dimensional_vector_space_pair begin + +definition "Basis_pair = B1 \ {0} \ {0} \ B2" + +sublocale p: finite_dimensional_vector_space scale Basis_pair +proof unfold_locales + show "finite Basis_pair" + by (auto intro!: finite_cartesian_product vs1.finite_Basis vs2.finite_Basis simp: Basis_pair_def) + show "p.independent Basis_pair" + unfolding p.dependent_def Basis_pair_def + proof safe + fix a + assume a: "a \ B1" + assume "(a, 0) \ p.span (B1 \ {0} \ {0} \ B2 - {(a, 0)})" + also have "B1 \ {0} \ {0} \ B2 - {(a, 0)} = (B1 - {a}) \ {0} \ {0} \ B2" + by auto + finally show False + using a vs1.dependent_def vs1.independent_Basis + by (auto simp: p.span_Un span_Times_sing1 span_Times_sing2) + next + fix b + assume b: "b \ B2" + assume "(0, b) \ p.span (B1 \ {0} \ {0} \ B2 - {(0, b)})" + also have "(B1 \ {0} \ {0} \ B2 - {(0, b)}) = B1 \ {0} \ {0} \ (B2 - {b})" + by auto + finally show False + using b vs2.dependent_def vs2.independent_Basis + by (auto simp: p.span_Un span_Times_sing1 span_Times_sing2) + qed + show "p.span Basis_pair = UNIV" + by (auto simp: p.span_Un span_Times_sing2 span_Times_sing1 vs1.span_Basis vs2.span_Basis + Basis_pair_def) +qed + +lemma dim_Times: + assumes "vs1.subspace S" "vs2.subspace T" + shows "p.dim(S \ T) = vs1.dim S + vs2.dim T" +proof - + interpret p1: Vector_Spaces.linear s1 scale "(\x. (x, 0))" + by unfold_locales (auto simp: scale_def) + interpret pair1: finite_dimensional_vector_space_pair "( *a)" B1 scale Basis_pair + by unfold_locales + interpret p2: Vector_Spaces.linear s2 scale "(\x. (0, x))" + by unfold_locales (auto simp: scale_def) + interpret pair2: finite_dimensional_vector_space_pair "( *b)" B2 scale Basis_pair + by unfold_locales + have ss: "p.subspace ((\x. (x, 0)) ` S)" "p.subspace (Pair 0 ` T)" + by (rule p1.subspace_image p2.subspace_image assms)+ + have "p.dim(S \ T) = p.dim({u + v |u v. u \ (\x. (x, 0)) ` S \ v \ Pair 0 ` T})" + by (simp add: Times_eq_image_sum) + moreover have "p.dim ((\x. (x, 0::'c)) ` S) = vs1.dim S" "p.dim (Pair (0::'b) ` T) = vs2.dim T" + by (simp_all add: inj_on_def p1.linear_axioms pair1.dim_image_eq p2.linear_axioms pair2.dim_image_eq) + moreover have "p.dim ((\x. (x, 0)) ` S \ Pair 0 ` T) = 0" + by (subst p.dim_eq_0) auto + ultimately show ?thesis + using p.dim_sums_Int [OF ss] by linarith +qed + +lemma dimension_pair: "p.dimension = vs1.dimension + vs2.dimension" + using dim_Times[OF vs1.subspace_UNIV vs2.subspace_UNIV] + by (auto simp: p.dim_UNIV vs1.dim_UNIV vs2.dim_UNIV + p.dimension_def vs1.dimension_def vs2.dimension_def) + end + +end diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Sigma_Algebra.thy --- a/src/HOL/Analysis/Sigma_Algebra.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Sigma_Algebra.thy Wed May 02 13:49:38 2018 +0200 @@ -11,7 +11,6 @@ imports Complex_Main "HOL-Library.Countable_Set" - "HOL-Library.FuncSet" "HOL-Library.Indicator_Function" "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Disjoint_Sets" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Starlike.thy Wed May 02 13:49:38 2018 +0200 @@ -10,7 +10,6 @@ theory Starlike imports Convex_Euclidean_Space - begin subsection \Midpoint\ @@ -108,8 +107,12 @@ using less_eq_real_def by (auto simp: segment algebra_simps) lemma closed_segment_linear_image: - "linear f \ closed_segment (f a) (f b) = f ` (closed_segment a b)" - by (force simp add: in_segment linear_add_cmul) + "closed_segment (f a) (f b) = f ` (closed_segment a b)" if "linear f" +proof - + interpret linear f by fact + show ?thesis + by (force simp add: in_segment add scale) +qed lemma open_segment_linear_image: "\linear f; inj f\ \ open_segment (f a) (f b) = f ` (open_segment a b)" @@ -1680,9 +1683,9 @@ fix x :: "'a::euclidean_space" assume "x \ d" then have "x \ span d" - using span_superset[of _ "d"] by auto + using span_base[of _ "d"] by auto then have "x /\<^sub>R (2 * real (card d)) \ span d" - using span_mul[of x "d" "(inverse (real (card d)) / 2)"] by auto + using span_scale[of x "d" "(inverse (real (card d)) / 2)"] by auto } then show "\x. x\d \ x /\<^sub>R (2 * real (card d)) \ span d" by auto @@ -1705,22 +1708,26 @@ then show ?thesis using rel_interior_sing by auto next case False - obtain B where B: "independent B \ B \ S \ S \ span B \ card B = dim S" + obtain B where B: "independent B \ B \ S \ S \ span B \ + card B = dim S" using basis_exists[of S] by auto then have "B \ {}" using B assms \S \ {0}\ span_empty by auto have "insert 0 B \ span B" - using subspace_span[of B] subspace_0[of "span B"] span_inc by auto + using subspace_span[of B] subspace_0[of "span B"] + span_superset by auto then have "span (insert 0 B) \ span B" using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast then have "convex hull insert 0 B \ span B" using convex_hull_subset_span[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) \ span B" - using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast + using span_span[of B] + span_mono[of "convex hull insert 0 B" "span B"] by blast then have *: "span (convex hull insert 0 B) = span B" using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) = span S" - using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto + using B span_mono[of B S] span_mono[of S "span B"] + span_span[of B] by auto moreover have "0 \ affine hull (convex hull insert 0 B)" using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" @@ -1787,8 +1794,8 @@ obtains a where "a \ interior (convex hull (insert 0 S))" proof - have "affine hull (insert 0 S) = UNIV" - apply (simp add: hull_inc affine_hull_span_0) - using assms dim_eq_full indep_card_eq_dim_span by fastforce + by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric] + assms(1) assms(3) dim_eq_card_independent) moreover have "rel_interior (convex hull insert 0 S) \ {}" using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto ultimately have "interior (convex hull insert 0 S) \ {}" @@ -2815,6 +2822,7 @@ using assms rel_interior_empty rel_interior_eq_empty by auto next case False + interpret linear f by fact have *: "f ` (rel_interior S) \ f ` S" unfolding image_mono using rel_interior_subset by auto have "f ` S \ f ` (closure S)" @@ -2847,15 +2855,15 @@ then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" - using x1 z1 \linear f\ by (simp add: linear_add_cmul) + using x1 z1 by (simp add: linear_add linear_scale \linear f\) ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" - using convex_rel_interior_iff[of "f ` S" z] \convex S\ - \linear f\ \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] + using convex_rel_interior_iff[of "f ` S" z] \convex S\ \linear f\ + \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] by auto } ultimately show ?thesis by auto @@ -2868,6 +2876,7 @@ and "f -` (rel_interior S) \ {}" shows "rel_interior (f -` S) = f -` (rel_interior S)" proof - + interpret linear f by fact have "S \ {}" using assms rel_interior_empty by auto have nonemp: "f -` S \ {}" @@ -2877,7 +2886,7 @@ have conv: "convex (f -` S)" using convex_linear_vimage assms by auto then have "convex (S \ range f)" - by (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image) + by (simp add: assms(2) convex_Int convex_linear_image linear_axioms) { fix z assume "z \ f -` (rel_interior S)" @@ -2917,7 +2926,7 @@ convex_rel_interior_iff[of "S \ (range f)" "f z"] by auto moreover have "affine (range f)" - by (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image) + by (simp add: linear_axioms linear_subspace_image subspace_imp_affine) ultimately have "f z \ rel_interior S" using convex_affine_rel_interior_Int[of S "range f"] assms by auto then have "z \ f -` (rel_interior S)" @@ -3888,9 +3897,7 @@ apply (induct s rule: finite_induct) apply (simp_all add: affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation) apply (rule empty_interior_lowdim) - apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) - apply (metis Suc_le_lessD not_less order_trans card_image_le finite_imageI dim_le_card) - done + by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans]) lemma empty_interior_convex_hull: fixes s :: "'a::euclidean_space set" @@ -5641,7 +5648,7 @@ by (simp add: algebra_simps) have "inverse k *\<^sub>R k *\<^sub>R (x-a) \ span ((\x. x - a) ` (S \ T))" apply (rule span_mul) - apply (rule span_superset) + apply (rule span_base) apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"]) apply (auto simp: S T) done @@ -5799,7 +5806,7 @@ by (auto simp: sum.remove [of _ k] inner_commute assms that) finally have "x = (\b\Basis - {k}. (x \ b) *\<^sub>R b)" . then show ?thesis - by (simp add: Linear_Algebra.span_finite) metis + by (simp add: span_finite) qed show ?thesis apply (rule span_subspace [symmetric]) @@ -5811,11 +5818,11 @@ lemma dim_special_hyperplane: fixes k :: "'n::euclidean_space" shows "k \ Basis \ dim {x. k \ x = 0} = DIM('n) - 1" -apply (simp add: special_hyperplane_span) -apply (rule Linear_Algebra.dim_unique [OF subset_refl]) -apply (auto simp: Diff_subset independent_substdbasis) -apply (metis member_remove remove_def span_clauses(1)) -done + apply (simp add: special_hyperplane_span) + apply (rule dim_unique [OF order_refl]) + apply (auto simp: Diff_subset independent_substdbasis) + apply (metis member_remove remove_def span_clauses(1)) + done proposition dim_hyperplane: fixes a :: "'a::euclidean_space" @@ -5831,7 +5838,7 @@ and ortho: "pairwise orthogonal B" using orthogonal_basis_exists by metis with assms have "a \ span B" - by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0 span_subspace) + by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0) then have ind: "independent (insert a B)" by (simp add: \independent B\ independent_insert) have "finite B" @@ -5843,13 +5850,15 @@ using assms by (auto simp: algebra_simps) show "y \ span (insert a B)" - by (metis (mono_tags, lifting) z Bsub Convex_Euclidean_Space.span_eq + by (metis (mono_tags, lifting) z Bsub span_eq_iff add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) qed then have dima: "DIM('a) = dim(insert a B)" - by (metis antisym dim_UNIV dim_subset_UNIV subset_le_dim) + by (metis independent_Basis span_Basis dim_eq_card top.extremum_uniqueI) then show ?thesis - by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \a \ span B\ ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB) + by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \a \ span B\ ind card0 + card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE + subspB) qed lemma lowdim_eq_hyperplane: @@ -5857,13 +5866,15 @@ assumes "dim S = DIM('a) - 1" obtains a where "a \ 0" and "span S = {x. a \ x = 0}" proof - - have [simp]: "dim S < DIM('a)" - by (simp add: DIM_positive assms) + have dimS: "dim S < DIM('a)" + by (simp add: assms) then obtain b where b: "b \ 0" "span S \ {a. b \ a = 0}" using lowdim_subset_hyperplane [of S] by fastforce show ?thesis - using b that real_vector_class.subspace_span [of S] - by (simp add: assms dim_hyperplane subspace_dim_equal subspace_hyperplane) + apply (rule that[OF b(1)]) + apply (rule subspace_dim_equal) + by (auto simp: assms b dim_hyperplane dim_span subspace_hyperplane + subspace_span) qed lemma dim_eq_hyperplane: @@ -5927,47 +5938,6 @@ subsection%unimportant\Some stepping theorems\ -lemma dim_empty [simp]: "dim ({} :: 'a::euclidean_space set) = 0" - by (force intro!: dim_unique) - -lemma dim_insert: - fixes x :: "'a::euclidean_space" - shows "dim (insert x S) = (if x \ span S then dim S else dim S + 1)" -proof - - show ?thesis - proof (cases "x \ span S") - case True then show ?thesis - by (metis dim_span span_redundant) - next - case False - obtain B where B: "B \ span S" "independent B" "span S \ span B" "card B = dim (span S)" - using basis_exists [of "span S"] by blast - have 1: "insert x B \ span (insert x S)" - by (meson \B \ span S\ dual_order.trans insertI1 insert_subsetI span_mono span_superset subset_insertI) - have 2: "span (insert x S) \ span (insert x B)" - by (metis \B \ span S\ \span S \ span B\ span_breakdown_eq span_subspace subsetI subspace_span) - have 3: "independent (insert x B)" - by (metis B independent_insert span_subspace subspace_span False) - have "dim (span (insert x S)) = Suc (dim S)" - apply (rule dim_unique [OF 1 2 3]) - by (metis B False card_insert_disjoint dim_span independent_imp_finite subsetCE) - then show ?thesis - by (simp add: False) - qed -qed - -lemma dim_singleton [simp]: - fixes x :: "'a::euclidean_space" - shows "dim{x} = (if x = 0 then 0 else 1)" -by (simp add: dim_insert) - -lemma dim_eq_0 [simp]: - fixes S :: "'a::euclidean_space set" - shows "dim S = 0 \ S \ {0}" -apply safe -apply (metis DIM_positive DIM_real card_ge_dim_independent contra_subsetD dim_empty dim_insert dim_singleton empty_subsetI independent_empty less_not_refl zero_le) -by (metis dim_singleton dim_subset le_0_eq) - lemma aff_dim_insert: fixes a :: "'a::euclidean_space" shows "aff_dim (insert a S) = (if a \ affine hull S then aff_dim S else aff_dim S + 1)" @@ -5981,9 +5951,7 @@ show ?thesis using S apply (simp add: hull_redundant cong: aff_dim_affine_hull2) apply (simp add: affine_hull_insert_span_gen hull_inc) - apply (simp add: insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert) - apply (metis (no_types, lifting) add_minus_cancel image_iff uminus_add_conv_diff) - done + by (force simp add:span_zero insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert) qed lemma affine_dependent_choose: @@ -6147,7 +6115,7 @@ then have "norm (a /\<^sub>R (norm a)) = 1" by simp moreover have "a /\<^sub>R (norm a) \ span S" - by (simp add: \a \ S\ span_mul span_superset) + by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp show ?thesis @@ -6171,7 +6139,7 @@ moreover have "norm (a /\<^sub>R (norm a)) = 1" using \a \ 0\ by simp moreover have "a /\<^sub>R (norm a) \ span S" - by (simp add: \a \ S\ span_mul span_superset) + by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp have aa: "a /\<^sub>R (norm a) \ (\c\C. {x. 0 \ c \ x})" @@ -6636,41 +6604,6 @@ qed qed -subsection%unimportant\Misc results about span\ - -lemma eq_span_insert_eq: - assumes "(x - y) \ span S" - shows "span(insert x S) = span(insert y S)" -proof - - have *: "span(insert x S) \ span(insert y S)" if "(x - y) \ span S" for x y - proof - - have 1: "(r *\<^sub>R x - r *\<^sub>R y) \ span S" for r - by (metis real_vector.scale_right_diff_distrib span_mul that) - have 2: "(z - k *\<^sub>R y) - k *\<^sub>R (x - y) = z - k *\<^sub>R x" for z k - by (simp add: real_vector.scale_right_diff_distrib) - show ?thesis - apply (clarsimp simp add: span_breakdown_eq) - by (metis 1 2 diff_add_cancel real_vector.scale_right_diff_distrib span_add_eq) - qed - show ?thesis - apply (intro subset_antisym * assms) - using assms subspace_neg subspace_span minus_diff_eq by force -qed - -lemma dim_psubset: - fixes S :: "'a :: euclidean_space set" - shows "span S \ span T \ dim S < dim T" -by (metis (no_types, hide_lams) dim_span less_le not_le subspace_dim_equal subspace_span) - - -lemma basis_subspace_exists: - fixes S :: "'a::euclidean_space set" - shows - "subspace S - \ \b. finite b \ b \ S \ - independent b \ span b = S \ card b = dim S" -by (metis span_subspace basis_exists independent_imp_finite) - lemma affine_hyperplane_sums_eq_UNIV_0: fixes S :: "'a :: euclidean_space set" assumes "affine S" @@ -6687,7 +6620,7 @@ using \a \ w \ 0\ span_induct subspace_hyperplane by auto moreover have "w \ span {x + y |x y. x \ S \ a \ y = 0}" using \w \ S\ - by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_superset) + by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_base) ultimately have span2: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" by blast have "a \ 0" using assms inner_zero_left by blast @@ -6704,7 +6637,7 @@ using DIM_lt apply simp done ultimately show ?thesis - by (simp add: subs) (metis (lifting) span_eq subs) + by (simp add: subs) (metis (lifting) span_eq_iff subs) qed proposition affine_hyperplane_sums_eq_UNIV: @@ -6751,108 +6684,6 @@ finally show ?thesis . qed -proposition dim_sums_Int: - fixes S :: "'a :: euclidean_space set" - assumes "subspace S" "subspace T" - shows "dim {x + y |x y. x \ S \ y \ T} + dim(S \ T) = dim S + dim T" -proof - - obtain B where B: "B \ S \ T" "S \ T \ span B" - and indB: "independent B" - and cardB: "card B = dim (S \ T)" - using basis_exists by blast - then obtain C D where "B \ C" "C \ S" "independent C" "S \ span C" - and "B \ D" "D \ T" "independent D" "T \ span D" - using maximal_independent_subset_extend - by (metis Int_subset_iff \B \ S \ T\ indB) - then have "finite B" "finite C" "finite D" - by (simp_all add: independent_imp_finite indB independent_bound) - have Beq: "B = C \ D" - apply (rule sym) - apply (rule spanning_subset_independent) - using \B \ C\ \B \ D\ apply blast - apply (meson \independent C\ independent_mono inf.cobounded1) - using B \C \ S\ \D \ T\ apply auto - done - then have Deq: "D = B \ (D - C)" - by blast - have CUD: "C \ D \ {x + y |x y. x \ S \ y \ T}" - apply safe - apply (metis add.right_neutral subsetCE \C \ S\ \subspace T\ set_eq_subset span_0 span_minimal) - apply (metis add.left_neutral subsetCE \D \ T\ \subspace S\ set_eq_subset span_0 span_minimal) - done - have "a v = 0" if 0: "(\v\C. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v) = 0" - and v: "v \ C \ (D-C)" for a v - proof - - have eq: "(\v\D - C. a v *\<^sub>R v) = - (\v\C. a v *\<^sub>R v)" - using that add_eq_0_iff by blast - have "(\v\D - C. a v *\<^sub>R v) \ S" - apply (subst eq) - apply (rule subspace_neg [OF \subspace S\]) - apply (rule subspace_sum [OF \subspace S\]) - by (meson subsetCE subspace_mul \C \ S\ \subspace S\) - moreover have "(\v\D - C. a v *\<^sub>R v) \ T" - apply (rule subspace_sum [OF \subspace T\]) - by (meson DiffD1 \D \ T\ \subspace T\ subset_eq subspace_def) - ultimately have "(\v \ D-C. a v *\<^sub>R v) \ span B" - using B by blast - then obtain e where e: "(\v\B. e v *\<^sub>R v) = (\v \ D-C. a v *\<^sub>R v)" - using span_finite [OF \finite B\] by blast - have "\c v. \(\v\C. c v *\<^sub>R v) = 0; v \ C\ \ c v = 0" - using independent_explicit \independent C\ by blast - define cc where "cc x = (if x \ B then a x + e x else a x)" for x - have [simp]: "C \ B = B" "D \ B = B" "C \ - B = C-D" "B \ (D - C) = {}" - using \B \ C\ \B \ D\ Beq by blast+ - have f2: "(\v\C \ D. e v *\<^sub>R v) = (\v\D - C. a v *\<^sub>R v)" - using Beq e by presburger - have f3: "(\v\C \ D. a v *\<^sub>R v) = (\v\C - D. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v) + (\v\C \ D. a v *\<^sub>R v)" - using \finite C\ \finite D\ sum.union_diff2 by blast - have f4: "(\v\C \ (D - C). a v *\<^sub>R v) = (\v\C. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v)" - by (meson Diff_disjoint \finite C\ \finite D\ finite_Diff sum.union_disjoint) - have "(\v\C. cc v *\<^sub>R v) = 0" - using 0 f2 f3 f4 - apply (simp add: cc_def Beq if_smult \finite C\ sum.If_cases algebra_simps sum.distrib) - apply (simp add: add.commute add.left_commute diff_eq) - done - then have "\v. v \ C \ cc v = 0" - using independent_explicit \independent C\ by blast - then have C0: "\v. v \ C - B \ a v = 0" - by (simp add: cc_def Beq) meson - then have [simp]: "(\x\C - B. a x *\<^sub>R x) = 0" - by simp - have "(\x\C. a x *\<^sub>R x) = (\x\B. a x *\<^sub>R x)" - proof - - have "C - D = C - B" - using Beq by blast - then show ?thesis - using Beq \(\x\C - B. a x *\<^sub>R x) = 0\ f3 f4 by auto - qed - with 0 have Dcc0: "(\v\D. a v *\<^sub>R v) = 0" - apply (subst Deq) - by (simp add: \finite B\ \finite D\ sum_Un) - then have D0: "\v. v \ D \ a v = 0" - using independent_explicit \independent D\ by blast - show ?thesis - using v C0 D0 Beq by blast - qed - then have "independent (C \ (D - C))" - by (simp add: independent_explicit \finite C\ \finite D\ sum_Un del: Un_Diff_cancel) - then have indCUD: "independent (C \ D)" by simp - have "dim (S \ T) = card B" - by (rule dim_unique [OF B indB refl]) - moreover have "dim S = card C" - by (metis \C \ S\ \independent C\ \S \ span C\ basis_card_eq_dim) - moreover have "dim T = card D" - by (metis \D \ T\ \independent D\ \T \ span D\ basis_card_eq_dim) - moreover have "dim {x + y |x y. x \ S \ y \ T} = card(C \ D)" - apply (rule dim_unique [OF CUD _ indCUD refl], clarify) - apply (meson \S \ span C\ \T \ span D\ span_add span_inc span_minimal subsetCE subspace_span sup.bounded_iff) - done - ultimately show ?thesis - using \B = C \ D\ [symmetric] - by (simp add: \independent C\ \independent D\ card_Un_Int independent_finite) -qed - - lemma aff_dim_sums_Int_0: assumes "affine S" and "affine T" @@ -6921,62 +6752,8 @@ by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) -lemma dim_Times: - fixes S :: "'a :: euclidean_space set" and T :: "'a set" - assumes "subspace S" "subspace T" - shows "dim(S \ T) = dim S + dim T" -proof - - have ss: "subspace ((\x. (x, 0)) ` S)" "subspace (Pair 0 ` T)" - by (rule subspace_linear_image, unfold_locales, auto simp: assms)+ - have "dim(S \ T) = dim({u + v |u v. u \ (\x. (x, 0)) ` S \ v \ Pair 0 ` T})" - by (simp add: Times_eq_image_sum) - moreover have "dim ((\x. (x, 0::'a)) ` S) = dim S" "dim (Pair (0::'a) ` T) = dim T" - by (auto simp: additive.intro linear.intro linear_axioms.intro inj_on_def intro: dim_image_eq) - moreover have "dim ((\x. (x, 0)) ` S \ Pair 0 ` T) = 0" - by (subst dim_eq_0) (force simp: zero_prod_def) - ultimately show ?thesis - using dim_sums_Int [OF ss] by linarith -qed - subsection\ Orthogonal bases, Gram-Schmidt process, and related theorems\ -lemma span_delete_0 [simp]: "span(S - {0}) = span S" -proof - show "span (S - {0}) \ span S" - by (blast intro!: span_mono) -next - have "span S \ span(insert 0 (S - {0}))" - by (blast intro!: span_mono) - also have "... \ span(S - {0})" - using span_insert_0 by blast - finally show "span S \ span (S - {0})" . -qed - -lemma span_image_scale: - assumes "finite S" and nz: "\x. x \ S \ c x \ 0" - shows "span ((\x. c x *\<^sub>R x) ` S) = span S" -using assms -proof (induction S arbitrary: c) - case (empty c) show ?case by simp -next - case (insert x F c) - show ?case - proof (intro set_eqI iffI) - fix y - assume "y \ span ((\x. c x *\<^sub>R x) ` insert x F)" - then show "y \ span (insert x F)" - using insert by (force simp: span_breakdown_eq) - next - fix y - assume "y \ span (insert x F)" - then show "y \ span ((\x. c x *\<^sub>R x) ` insert x F)" - using insert - apply (clarsimp simp: span_breakdown_eq) - apply (rule_tac x="k / c x" in exI) - by simp - qed -qed - lemma pairwise_orthogonal_independent: assumes "pairwise orthogonal S" and "0 \ S" shows "independent S" @@ -7022,9 +6799,8 @@ lemma orthogonal_to_span: assumes a: "a \ span S" and x: "\y. y \ S \ orthogonal x y" shows "orthogonal x a" -apply (rule span_induct [OF a subspace_orthogonal_to_vector]) -apply (simp add: x) -done + by (metis a orthogonal_clauses(1,2,4) + span_induct_alt x) proposition%important Gram_Schmidt_step: fixes S :: "'a::euclidean_space set" @@ -7064,19 +6840,20 @@ define a' where "a' = a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)" obtain U where orthU: "pairwise orthogonal (S \ insert a' U)" and spanU: "span (insert a' S \ U) = span (insert a' S \ T)" - apply (rule exE [OF insert.IH [of "insert a' S"]]) - apply (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) - done + by (rule exE [OF insert.IH [of "insert a' S"]]) + (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute + pairwise_orthogonal_insert span_clauses) have orthS: "\x. x \ S \ a' \ x = 0" apply (simp add: a'_def) using Gram_Schmidt_step [OF \pairwise orthogonal S\] - apply (force simp: orthogonal_def inner_commute span_inc [THEN subsetD]) + apply (force simp: orthogonal_def inner_commute span_superset [THEN subsetD]) done have "span (S \ insert a' U) = span (insert a' (S \ T))" using spanU by simp also have "... = span (insert a (S \ T))" apply (rule eq_span_insert_eq) - apply (simp add: a'_def span_neg span_sum span_clauses(1) span_mul) + apply (simp add: a'_def span_neg span_sum span_clauses(1) + span_scale) done also have "... = span (S \ insert a T)" by simp @@ -7097,17 +6874,15 @@ with orthogonal_extension_aux [of B S] obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ B)" using assms pairwise_orthogonal_imp_finite by auto - show ?thesis - apply (rule_tac U=U in that) - apply (simp add: \pairwise orthogonal (S \ U)\) - by (metis \span (S \ U) = span (S \ B)\ \span B = span T\ span_Un) + with \span B = span T\ show ?thesis + by (rule_tac U=U in that) (auto simp: span_Un) qed corollary orthogonal_extension_strong: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" obtains U where "U \ (insert 0 S) = {}" "pairwise orthogonal (S \ U)" - "span (S \ U) = span (S \ T)" + "span (S \ U) = span (S \ T)" proof - obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" using orthogonal_extension assms by blast @@ -7115,8 +6890,8 @@ apply (rule_tac U = "U - (insert 0 S)" in that) apply blast apply (force simp: pairwise_def) - apply (metis (no_types, lifting) Un_Diff_cancel span_insert_0 span_Un) - done + apply (metis Un_Diff_cancel Un_insert_left span_redundant span_zero) + done qed subsection\Decomposing a vector into parts in orthogonal subspaces\ @@ -7132,7 +6907,7 @@ using basis_exists by blast with orthogonal_extension [of "{}" B] show ?thesis - by (metis Un_empty_left assms pairwise_empty span_inc span_subspace that) + by (metis Un_empty_left assms pairwise_empty span_superset span_subspace that) qed lemma orthogonal_basis_subspace: @@ -7161,7 +6936,7 @@ and "independent B" "card B = dim S" "span B = S" by (blast intro: orthogonal_basis_subspace [OF assms]) have 1: "(\x. x /\<^sub>R norm x) ` B \ S" - using \span B = S\ span_clauses(1) span_mul by fastforce + using \span B = S\ span_clauses(1) span_scale by fastforce have 2: "pairwise orthogonal ((\x. x /\<^sub>R norm x) ` B)" using orth by (force simp: pairwise_def orthogonal_clauses) have 3: "\x. x \ (\x. x /\<^sub>R norm x) ` B \ norm x = 1" @@ -7194,7 +6969,8 @@ obtain B where "B \ span S" and orthB: "pairwise orthogonal B" and "\x. x \ B \ norm x = 1" and "independent B" "card B = dim S" "span B = span S" - by (rule orthonormal_basis_subspace [of "span S"]) auto + by (rule orthonormal_basis_subspace [of "span S", OF subspace_span]) + (auto simp: dim_span) with assms obtain u where spanBT: "span B \ span T" and "u \ span B" "u \ span T" by auto obtain C where orthBC: "pairwise orthogonal (B \ C)" and spanBC: "span (B \ C) = span (B \ {u})" @@ -7203,18 +6979,21 @@ proof (cases "C \ insert 0 B") case True then have "C \ span B" - using Linear_Algebra.span_eq + using span_eq by (metis span_insert_0 subset_trans) moreover have "u \ span (B \ C)" - using \span (B \ C) = span (B \ {u})\ span_inc by force + using \span (B \ C) = span (B \ {u})\ span_superset by force ultimately show ?thesis - by (metis \u \ span B\ span_Un span_span sup.orderE) + using True \u \ span B\ + by (metis Un_insert_left span_insert_0 sup.orderE) next case False then obtain x where "x \ C" "x \ 0" "x \ B" by blast then have "x \ span T" - by (metis (no_types, lifting) Un_insert_right Un_upper2 \u \ span T\ spanBT spanBC \u \ span T\ insert_subset span_inc span_mono span_span subsetCE subset_trans sup_bot.comm_neutral) + by (metis (no_types, lifting) Un_insert_right Un_upper2 \u \ span T\ spanBT spanBC + \u \ span T\ insert_subset span_superset span_mono + span_span subsetCE subset_trans sup_bot.comm_neutral) moreover have "orthogonal x y" if "y \ span B" for y using that proof (rule span_induct) @@ -7234,8 +7013,10 @@ obtains x where "x \ 0" "\y. y \ span S \ orthogonal x y" proof - have "span S \ UNIV" - by (metis assms dim_eq_full less_irrefl top.not_eq_extremum) - with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis by auto + by (metis (mono_tags) UNIV_I assms inner_eq_zero_iff less_le lowdim_subset_hyperplane + mem_Collect_eq top.extremum_strict top.not_eq_extremum) + with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis + by (auto simp: span_UNIV) qed corollary orthogonal_to_vector_exists: @@ -7253,14 +7034,16 @@ fixes S :: "'a :: euclidean_space set" obtains y z where "y \ span S" "\w. w \ span S \ orthogonal z w" "x = y + z" proof%unimportant - - obtain T where "0 \ T" "T \ span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S" + obtain T where "0 \ T" "T \ span S" "pairwise orthogonal T" "independent T" + "card T = dim (span S)" "span T = span S" using orthogonal_basis_subspace subspace_span by blast let ?a = "\b\T. (b \ x / (b \ b)) *\<^sub>R b" have orth: "orthogonal (x - ?a) w" if "w \ span S" for w - by (simp add: Gram_Schmidt_step \pairwise orthogonal T\ \span T = span S\ orthogonal_commute that) + by (simp add: Gram_Schmidt_step \pairwise orthogonal T\ \span T = span S\ + orthogonal_commute that) show ?thesis apply (rule_tac y = "?a" and z = "x - ?a" in that) - apply (meson \T \ span S\ span_mul span_sum subsetCE) + apply (meson \T \ span S\ span_scale span_sum subsetCE) apply (fact orth, simp) done qed @@ -7284,7 +7067,8 @@ lemma vector_in_orthogonal_spanningset: fixes a :: "'a::euclidean_space" obtains S where "a \ S" "pairwise orthogonal S" "span S = UNIV" - by (metis UNIV_I Un_iff empty_iff insert_subset orthogonal_extension pairwise_def pairwise_orthogonal_insert span_UNIV subsetI subset_antisym) + by (metis UNIV_I Un_iff empty_iff insert_subset orthogonal_extension pairwise_def + pairwise_orthogonal_insert span_UNIV subsetI subset_antisym) lemma vector_in_orthogonal_basis: fixes a :: "'a::euclidean_space" @@ -7304,7 +7088,7 @@ using \independent (S - {0})\ independent_finite by blast show "card (S - {0}) = DIM('a)" using span_delete_0 [of S] S - by (simp add: \independent (S - {0})\ indep_card_eq_dim_span) + by (simp add: \independent (S - {0})\ indep_card_eq_dim_span dim_UNIV) qed (use S \a \ 0\ in auto) qed @@ -7337,7 +7121,9 @@ then show "card ?S = DIM('a)" by (simp add: card_image S) show "span ?S = UNIV" - by (metis (no_types) \0 \ S\ \finite S\ \span S = UNIV\ field_class.field_inverse_zero inverse_inverse_eq less_irrefl span_image_scale zero_less_norm_iff) + by (metis (no_types) \0 \ S\ \finite S\ \span S = UNIV\ + field_class.field_inverse_zero inverse_inverse_eq less_irrefl span_image_scale + zero_less_norm_iff) qed qed @@ -7354,17 +7140,17 @@ then have 0: "\x y. \x \ span A; y \ span B\ \ x \ y = 0" by simp have "dim(A \ B) = dim (span (A \ B))" - by simp - also have "... = dim ((\(a, b). a + b) ` (span A \ span B))" - by (simp add: span_Un) - also have "... = dim {x + y |x y. x \ span A \ y \ span B}" + by (simp add: dim_span) + also have "span (A \ B) = ((\(a, b). a + b) ` (span A \ span B))" + by (auto simp add: span_Un image_def) + also have "dim \ = dim {x + y |x y. x \ span A \ y \ span B}" by (auto intro!: arg_cong [where f=dim]) also have "... = dim {x + y |x y. x \ span A \ y \ span B} + dim(span A \ span B)" by (auto simp: dest: 0) also have "... = dim (span A) + dim (span B)" - by (rule dim_sums_Int) auto + by (rule dim_sums_Int) (auto simp: subspace_span) also have "... = dim A + dim B" - by simp + by (simp add: dim_span) finally show ?thesis . qed @@ -7384,22 +7170,24 @@ obtain y z where "x = y + z" "y \ span A" and orth: "\w. w \ span A \ orthogonal z w" using orthogonal_subspace_decomp_exists [of A x] that by auto have "y \ span B" - by (metis span_eq \y \ span A\ assms subset_iff) + using \y \ span A\ assms(3) span_mono by blast then have "z \ {a \ B. \x. x \ A \ orthogonal x a}" - by simp (metis (no_types) span_eq \x = y + z\ \subspace A\ \subspace B\ orth orthogonal_commute span_add_eq that) + apply simp + using \x = y + z\ assms(1) assms(2) orth orthogonal_commute span_add_eq + span_eq_iff that by blast then have z: "z \ span {y \ B. \x\A. orthogonal x y}" - by (meson span_inc subset_iff) + by (meson span_superset subset_iff) then show ?thesis - apply (simp add: span_Un image_def) - apply (rule bexI [OF _ z]) - apply (simp add: \x = y + z\ \y \ span A\) - done + apply (auto simp: span_Un image_def \x = y + z\ \y \ span A\) + using \y \ span A\ add.commute by blast qed show "span B \ span ({y \ B. \x\A. orthogonal x y} \ A)" - by (rule span_minimal) (auto intro: * span_minimal elim: ) + by (rule span_minimal) + (auto intro: * span_minimal simp: subspace_span) qed then show ?thesis - by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def) + by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq + orthogonal_commute orthogonal_def) qed lemma aff_dim_openin: @@ -7444,9 +7232,9 @@ then show "(\x. e *\<^sub>R x) ` B \ (\x. x - a) ` S" using e' by blast show "independent ((\x. e *\<^sub>R x) ` B)" - using \independent B\ - apply (rule independent_injective_image, simp) - by (metis \0 < e\ injective_scaleR less_irrefl) + using linear_scale_self \independent B\ + apply (rule linear_independent_injective_image) + using \0 < e\ inj_on_def by fastforce qed also have "... = aff_dim S" using \a \ S\ aff_dim_eq_dim hull_inc by force @@ -7466,7 +7254,8 @@ using aff_dim_openin by (metis aff_dim_subspace \subspace T\ \S \ {}\ ope subspace_affine) also have "... \ dim S" - by (metis aff_dim_subset aff_dim_subspace dim_span span_inc subspace_span) + by (metis aff_dim_subset aff_dim_subspace dim_span span_superset + subspace_span) finally show "dim T \ dim S" by simp qed @@ -7502,13 +7291,13 @@ using \0 < e\ by (simp add: dist_norm) next have "?y \ S" - by (metis span_eq \a \ span S\ \x \ S\ \subspace S\ subspace_add subspace_mul) + by (metis \a \ span S\ \x \ S\ assms(2) span_eq_iff subspace_add subspace_scale) moreover have "?y \ U" proof - have "e/2 / norm a \ 0" using \0 < e\ \a \ 0\ by auto then show ?thesis - by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_clauses(1)) + by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self scale_eq_0_iff span_add_eq span_clauses(1)) qed ultimately show "?y \ S - U" by blast qed @@ -7628,7 +7417,7 @@ then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" - by (metis subspace_add a' span_eq sub) + by (metis subspace_add a' span_eq_iff sub) then have Sclo: "\w. w \ S \ (w + a') \ S" by fastforce show ?thesis @@ -7685,7 +7474,7 @@ also have "... = dim {x - c |x. x \ B}" by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) also have "... = dim {x - c | x. x \ affine hull B}" - by (simp add: diffs_affine_hull_span \c \ B\) + by (simp add: diffs_affine_hull_span \c \ B\ dim_span) also have "... = dim {x - a |x. x \ S}" by (simp add: affS aff *) finally show ?thesis . @@ -7730,7 +7519,7 @@ by (simp add: aff_dim_linear_image_le assms(1)) next obtain g where "linear g" "g \ f = id" - using linear_injective_left_inverse assms by blast + using assms(1) assms(2) linear_injective_left_inverse by blast then have "aff_dim S \ aff_dim(g ` f ` S)" by (simp add: image_comp) also have "... \ aff_dim (f ` S)" @@ -7739,43 +7528,6 @@ qed -text\Choosing a subspace of a given dimension\ -proposition choose_subspace_of_subspace: - fixes S :: "'n::euclidean_space set" - assumes "n \ dim S" - obtains T where "subspace T" "T \ span S" "dim T = n" -proof - - have "\T. subspace T \ T \ span S \ dim T = n" - using assms - proof (induction n) - case 0 then show ?case by force - next - case (Suc n) - then obtain T where "subspace T" "T \ span S" "dim T = n" - by force - then show ?case - proof (cases "span S \ span T") - case True - have "dim S = dim T" - apply (rule span_eq_dim [OF subset_antisym [OF True]]) - by (simp add: \T \ span S\ span_minimal subspace_span) - then show ?thesis - using Suc.prems \dim T = n\ by linarith - next - case False - then obtain y where y: "y \ S" "y \ T" - by (meson span_mono subsetI) - then have "span (insert y T) \ span S" - by (metis (no_types) \T \ span S\ subsetD insert_subset span_inc span_mono span_span) - with \dim T = n\ \subspace T\ y show ?thesis - apply (rule_tac x="span(insert y T)" in exI) - apply (auto simp: dim_insert) - using span_eq by blast - qed - qed - with that show ?thesis by blast -qed - lemma choose_affine_subset: assumes "affine S" "-1 \ d" and dle: "d \ aff_dim S" obtains T where "affine T" "T \ S" "aff_dim T = d" @@ -7799,7 +7551,7 @@ moreover have "(+) a ` T \ S" proof - have "T \ (+) (- a) ` S" - by (metis (no_types) span_eq Tsb ss) + by (metis (no_types) span_eq_iff Tsb ss) then show "(+) a ` T \ S" using add_ac by auto qed @@ -8077,7 +7829,7 @@ also have "... = (\x \ S - {b}. v x *\<^sub>R x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x *\<^sub>R x)" - by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert real_vector.scale_eq_0_iff sum_diff1) + by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1) finally show "(\x\insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x" by (simp add: vx) qed @@ -8123,7 +7875,7 @@ apply (rule sum.cong, auto) done also have "... = (v b / u b) *\<^sub>R a + (\x \ S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\x \ S - {b}. u x *\<^sub>R x)" - by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left real_vector.scale_sum_right) + by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right) also have "... = (\x\S. v x *\<^sub>R x)" using \0 < u b\ True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps) finally @@ -8361,7 +8113,8 @@ have "v = ?u + (v - ?u)" by simp moreover have "?u \ U" - by (metis (no_types, lifting) \span B = U\ assms real_vector_class.subspace_sum span_clauses(1) span_mul) + by (metis (no_types, lifting) \span B = U\ assms subspace_sum + span_clauses(1) span_scale) moreover have "(v - ?u) \ U\<^sup>\" proof (clarsimp simp: orthogonal_comp_def orthogonal_def) fix y @@ -8381,33 +8134,6 @@ by auto qed -lemma add_subspaces: - assumes "subspace S" "subspace T" - shows "subspace (S + T)" - unfolding subspace_def -proof (intro conjI ballI allI) - show "0 \ S + T" - by (meson assms set_zero_plus2 subsetCE subspace_0) -next - fix x y - assume "x \ S + T" and "y \ S + T" - then obtain xs xt ys yt where "xs \ S" "xt \ T" "ys \ S" "yt \ T" and eq: "x = xs+xt" "y = ys+yt" - by (meson set_plus_elim) - then have "xs+ys \ S" "xt+yt \ T" - using assms subspace_def by blast+ - then have "(xs + ys) + (xt + yt) \ S + T" - by blast - then show "x + y \ S + T" - by (simp add: eq add.assoc add.left_commute) -next - fix c x - assume "x \ S + T" - then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" - by (meson set_plus_elim) - then show "c *\<^sub>R x \ S + T" - by (metis assms scaleR_add_right set_plus_intro subspace_def) -qed - lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}" @@ -8480,7 +8206,9 @@ ultimately have "range (adjoint f)\<^sup>\\<^sup>\ = UNIV" by (metis orthogonal_comp_null) then show "surj (adjoint f)" - by (simp add: adjoint_linear \linear f\ subspace_linear_image orthogonal_comp_self) + using adjoint_linear \linear f\ + by (subst (asm) orthogonal_comp_self) + (simp add: adjoint_linear linear_subspace_image) qed lemma inj_adjoint_iff_surj [simp]: @@ -8513,7 +8241,8 @@ assume "\inj f" then show ?rhs using all_zero_iff - by (metis (no_types, hide_lams) adjoint_clauses(2) adjoint_linear assms linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) + by (metis (no_types, hide_lams) adjoint_clauses(2) adjoint_linear assms + linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) next assume ?rhs then show "\inj f" @@ -8527,4 +8256,3 @@ using assms by (fastforce simp add: linear_singular_into_hyperplane) end - diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Wed May 02 13:49:38 2018 +0200 @@ -10,18 +10,12 @@ imports "HOL-Library.Indicator_Function" "HOL-Library.Countable_Set" - "HOL-Library.FuncSet" Linear_Algebra Norm_Arith begin (* FIXME: move elsewhere *) -lemma Times_eq_image_sum: - fixes S :: "'a :: comm_monoid_add set" and T :: "'b :: comm_monoid_add set" - shows "S \ T = {u + v |u v. u \ (\x. (x, 0)) ` S \ v \ Pair 0 ` T}" - by force - lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" "{x. b \ a \ x} \ {x. a \ x \ b} = {x. a \ x = b}" diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Analysis/Weierstrass_Theorems.thy --- a/src/HOL/Analysis/Weierstrass_Theorems.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Analysis/Weierstrass_Theorems.thy Wed May 02 13:49:38 2018 +0200 @@ -1376,7 +1376,7 @@ using \polynomial_function g\ by auto show "(\x. \i\B. (g x \ i) *\<^sub>R i) ` S \ T" using \B \ T\ - by (blast intro: real_vector_class.subspace_sum subspace_mul \subspace T\) + by (blast intro: subspace_sum subspace_mul \subspace T\) show "norm (f x - (\i\B. (g x \ i) *\<^sub>R i)) < e" if "x \ S" for x proof - have orth': "pairwise (\i j. orthogonal ((f x \ i) *\<^sub>R i - (g x \ i) *\<^sub>R i) diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Computational_Algebra/Formal_Power_Series.thy --- a/src/HOL/Computational_Algebra/Formal_Power_Series.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Computational_Algebra/Formal_Power_Series.thy Wed May 02 13:49:38 2018 +0200 @@ -105,9 +105,6 @@ lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)" by auto -lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)" - by auto - subsection \Formal power series form a commutative ring with unity, if the range of sequences they represent is a commutative ring with unity\ diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/FuncSet.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/FuncSet.thy Wed May 02 13:49:38 2018 +0200 @@ -0,0 +1,568 @@ +(* Title: HOL/FuncSet.thy + Author: Florian Kammueller and Lawrence C Paulson, Lukas Bulwahn +*) + +section \Pi and Function Sets\ + +theory FuncSet + imports Main + abbrevs PiE = "Pi\<^sub>E" + and PIE = "\\<^sub>E" +begin + +definition Pi :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" + where "Pi A B = {f. \x. x \ A \ f x \ B x}" + +definition extensional :: "'a set \ ('a \ 'b) set" + where "extensional A = {f. \x. x \ A \ f x = undefined}" + +definition "restrict" :: "('a \ 'b) \ 'a set \ 'a \ 'b" + where "restrict f A = (\x. if x \ A then f x else undefined)" + +abbreviation funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 60) + where "A \ B \ Pi A (\_. B)" + +syntax + "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\ _\_./ _)" 10) + "_lam" :: "pttrn \ 'a set \ ('a \ 'b) \ ('a \ 'b)" ("(3\_\_./ _)" [0,0,3] 3) +translations + "\ x\A. B" \ "CONST Pi A (\x. B)" + "\x\A. f" \ "CONST restrict (\x. f) A" + +definition "compose" :: "'a set \ ('b \ 'c) \ ('a \ 'b) \ ('a \ 'c)" + where "compose A g f = (\x\A. g (f x))" + + +subsection \Basic Properties of @{term Pi}\ + +lemma Pi_I[intro!]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" + by (simp add: Pi_def) + +lemma Pi_I'[simp]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" + by (simp add:Pi_def) + +lemma funcsetI: "(\x. x \ A \ f x \ B) \ f \ A \ B" + by (simp add: Pi_def) + +lemma Pi_mem: "f \ Pi A B \ x \ A \ f x \ B x" + by (simp add: Pi_def) + +lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" + unfolding Pi_def by auto + +lemma PiE [elim]: "f \ Pi A B \ (f x \ B x \ Q) \ (x \ A \ Q) \ Q" + by (auto simp: Pi_def) + +lemma Pi_cong: "(\w. w \ A \ f w = g w) \ f \ Pi A B \ g \ Pi A B" + by (auto simp: Pi_def) + +lemma funcset_id [simp]: "(\x. x) \ A \ A" + by auto + +lemma funcset_mem: "f \ A \ B \ x \ A \ f x \ B" + by (simp add: Pi_def) + +lemma funcset_image: "f \ A \ B \ f ` A \ B" + by auto + +lemma image_subset_iff_funcset: "F ` A \ B \ F \ A \ B" + by auto + +lemma Pi_eq_empty[simp]: "(\ x \ A. B x) = {} \ (\x\A. B x = {})" + apply (simp add: Pi_def) + apply auto + txt \Converse direction requires Axiom of Choice to exhibit a function + picking an element from each non-empty @{term "B x"}\ + apply (drule_tac x = "\u. SOME y. y \ B u" in spec) + apply auto + apply (cut_tac P = "\y. y \ B x" in some_eq_ex) + apply auto + done + +lemma Pi_empty [simp]: "Pi {} B = UNIV" + by (simp add: Pi_def) + +lemma Pi_Int: "Pi I E \ Pi I F = (\ i\I. E i \ F i)" + by auto + +lemma Pi_UN: + fixes A :: "nat \ 'i \ 'a set" + assumes "finite I" + and mono: "\i n m. i \ I \ n \ m \ A n i \ A m i" + shows "(\n. Pi I (A n)) = (\ i\I. \n. A n i)" +proof (intro set_eqI iffI) + fix f + assume "f \ (\ i\I. \n. A n i)" + then have "\i\I. \n. f i \ A n i" + by auto + from bchoice[OF this] obtain n where n: "f i \ A (n i) i" if "i \ I" for i + by auto + obtain k where k: "n i \ k" if "i \ I" for i + using \finite I\ finite_nat_set_iff_bounded_le[of "n`I"] by auto + have "f \ Pi I (A k)" + proof (intro Pi_I) + fix i + assume "i \ I" + from mono[OF this, of "n i" k] k[OF this] n[OF this] + show "f i \ A k i" by auto + qed + then show "f \ (\n. Pi I (A n))" + by auto +qed auto + +lemma Pi_UNIV [simp]: "A \ UNIV = UNIV" + by (simp add: Pi_def) + +text \Covariance of Pi-sets in their second argument\ +lemma Pi_mono: "(\x. x \ A \ B x \ C x) \ Pi A B \ Pi A C" + by auto + +text \Contravariance of Pi-sets in their first argument\ +lemma Pi_anti_mono: "A' \ A \ Pi A B \ Pi A' B" + by auto + +lemma prod_final: + assumes 1: "fst \ f \ Pi A B" + and 2: "snd \ f \ Pi A C" + shows "f \ (\ z \ A. B z \ C z)" +proof (rule Pi_I) + fix z + assume z: "z \ A" + have "f z = (fst (f z), snd (f z))" + by simp + also have "\ \ B z \ C z" + by (metis SigmaI PiE o_apply 1 2 z) + finally show "f z \ B z \ C z" . +qed + +lemma Pi_split_domain[simp]: "x \ Pi (I \ J) X \ x \ Pi I X \ x \ Pi J X" + by (auto simp: Pi_def) + +lemma Pi_split_insert_domain[simp]: "x \ Pi (insert i I) X \ x \ Pi I X \ x i \ X i" + by (auto simp: Pi_def) + +lemma Pi_cancel_fupd_range[simp]: "i \ I \ x \ Pi I (B(i := b)) \ x \ Pi I B" + by (auto simp: Pi_def) + +lemma Pi_cancel_fupd[simp]: "i \ I \ x(i := a) \ Pi I B \ x \ Pi I B" + by (auto simp: Pi_def) + +lemma Pi_fupd_iff: "i \ I \ f \ Pi I (B(i := A)) \ f \ Pi (I - {i}) B \ f i \ A" + apply auto + apply (drule_tac x=x in Pi_mem) + apply (simp_all split: if_split_asm) + apply (drule_tac x=i in Pi_mem) + apply (auto dest!: Pi_mem) + done + + +subsection \Composition With a Restricted Domain: @{term compose}\ + +lemma funcset_compose: "f \ A \ B \ g \ B \ C \ compose A g f \ A \ C" + by (simp add: Pi_def compose_def restrict_def) + +lemma compose_assoc: + assumes "f \ A \ B" + and "g \ B \ C" + and "h \ C \ D" + shows "compose A h (compose A g f) = compose A (compose B h g) f" + using assms by (simp add: fun_eq_iff Pi_def compose_def restrict_def) + +lemma compose_eq: "x \ A \ compose A g f x = g (f x)" + by (simp add: compose_def restrict_def) + +lemma surj_compose: "f ` A = B \ g ` B = C \ compose A g f ` A = C" + by (auto simp add: image_def compose_eq) + + +subsection \Bounded Abstraction: @{term restrict}\ + +lemma restrict_cong: "I = J \ (\i. i \ J =simp=> f i = g i) \ restrict f I = restrict g J" + by (auto simp: restrict_def fun_eq_iff simp_implies_def) + +lemma restrict_in_funcset: "(\x. x \ A \ f x \ B) \ (\x\A. f x) \ A \ B" + by (simp add: Pi_def restrict_def) + +lemma restrictI[intro!]: "(\x. x \ A \ f x \ B x) \ (\x\A. f x) \ Pi A B" + by (simp add: Pi_def restrict_def) + +lemma restrict_apply[simp]: "(\y\A. f y) x = (if x \ A then f x else undefined)" + by (simp add: restrict_def) + +lemma restrict_apply': "x \ A \ (\y\A. f y) x = f x" + by simp + +lemma restrict_ext: "(\x. x \ A \ f x = g x) \ (\x\A. f x) = (\x\A. g x)" + by (simp add: fun_eq_iff Pi_def restrict_def) + +lemma restrict_UNIV: "restrict f UNIV = f" + by (simp add: restrict_def) + +lemma inj_on_restrict_eq [simp]: "inj_on (restrict f A) A = inj_on f A" + by (simp add: inj_on_def restrict_def) + +lemma Id_compose: "f \ A \ B \ f \ extensional A \ compose A (\y\B. y) f = f" + by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) + +lemma compose_Id: "g \ A \ B \ g \ extensional A \ compose A g (\x\A. x) = g" + by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) + +lemma image_restrict_eq [simp]: "(restrict f A) ` A = f ` A" + by (auto simp add: restrict_def) + +lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \ B)" + unfolding restrict_def by (simp add: fun_eq_iff) + +lemma restrict_fupd[simp]: "i \ I \ restrict (f (i := x)) I = restrict f I" + by (auto simp: restrict_def) + +lemma restrict_upd[simp]: "i \ I \ (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" + by (auto simp: fun_eq_iff) + +lemma restrict_Pi_cancel: "restrict x I \ Pi I A \ x \ Pi I A" + by (auto simp: restrict_def Pi_def) + + +subsection \Bijections Between Sets\ + +text \The definition of @{const bij_betw} is in \Fun.thy\, but most of +the theorems belong here, or need at least @{term Hilbert_Choice}.\ + +lemma bij_betwI: + assumes "f \ A \ B" + and "g \ B \ A" + and g_f: "\x. x\A \ g (f x) = x" + and f_g: "\y. y\B \ f (g y) = y" + shows "bij_betw f A B" + unfolding bij_betw_def +proof + show "inj_on f A" + by (metis g_f inj_on_def) + have "f ` A \ B" + using \f \ A \ B\ by auto + moreover + have "B \ f ` A" + by auto (metis Pi_mem \g \ B \ A\ f_g image_iff) + ultimately show "f ` A = B" + by blast +qed + +lemma bij_betw_imp_funcset: "bij_betw f A B \ f \ A \ B" + by (auto simp add: bij_betw_def) + +lemma inj_on_compose: "bij_betw f A B \ inj_on g B \ inj_on (compose A g f) A" + by (auto simp add: bij_betw_def inj_on_def compose_eq) + +lemma bij_betw_compose: "bij_betw f A B \ bij_betw g B C \ bij_betw (compose A g f) A C" + apply (simp add: bij_betw_def compose_eq inj_on_compose) + apply (auto simp add: compose_def image_def) + done + +lemma bij_betw_restrict_eq [simp]: "bij_betw (restrict f A) A B = bij_betw f A B" + by (simp add: bij_betw_def) + + +subsection \Extensionality\ + +lemma extensional_empty[simp]: "extensional {} = {\x. undefined}" + unfolding extensional_def by auto + +lemma extensional_arb: "f \ extensional A \ x \ A \ f x = undefined" + by (simp add: extensional_def) + +lemma restrict_extensional [simp]: "restrict f A \ extensional A" + by (simp add: restrict_def extensional_def) + +lemma compose_extensional [simp]: "compose A f g \ extensional A" + by (simp add: compose_def) + +lemma extensionalityI: + assumes "f \ extensional A" + and "g \ extensional A" + and "\x. x \ A \ f x = g x" + shows "f = g" + using assms by (force simp add: fun_eq_iff extensional_def) + +lemma extensional_restrict: "f \ extensional A \ restrict f A = f" + by (rule extensionalityI[OF restrict_extensional]) auto + +lemma extensional_subset: "f \ extensional A \ A \ B \ f \ extensional B" + unfolding extensional_def by auto + +lemma inv_into_funcset: "f ` A = B \ (\x\B. inv_into A f x) \ B \ A" + by (unfold inv_into_def) (fast intro: someI2) + +lemma compose_inv_into_id: "bij_betw f A B \ compose A (\y\B. inv_into A f y) f = (\x\A. x)" + apply (simp add: bij_betw_def compose_def) + apply (rule restrict_ext, auto) + done + +lemma compose_id_inv_into: "f ` A = B \ compose B f (\y\B. inv_into A f y) = (\x\B. x)" + apply (simp add: compose_def) + apply (rule restrict_ext) + apply (simp add: f_inv_into_f) + done + +lemma extensional_insert[intro, simp]: + assumes "a \ extensional (insert i I)" + shows "a(i := b) \ extensional (insert i I)" + using assms unfolding extensional_def by auto + +lemma extensional_Int[simp]: "extensional I \ extensional I' = extensional (I \ I')" + unfolding extensional_def by auto + +lemma extensional_UNIV[simp]: "extensional UNIV = UNIV" + by (auto simp: extensional_def) + +lemma restrict_extensional_sub[intro]: "A \ B \ restrict f A \ extensional B" + unfolding restrict_def extensional_def by auto + +lemma extensional_insert_undefined[intro, simp]: + "a \ extensional (insert i I) \ a(i := undefined) \ extensional I" + unfolding extensional_def by auto + +lemma extensional_insert_cancel[intro, simp]: + "a \ extensional I \ a \ extensional (insert i I)" + unfolding extensional_def by auto + + +subsection \Cardinality\ + +lemma card_inj: "f \ A \ B \ inj_on f A \ finite B \ card A \ card B" + by (rule card_inj_on_le) auto + +lemma card_bij: + assumes "f \ A \ B" "inj_on f A" + and "g \ B \ A" "inj_on g B" + and "finite A" "finite B" + shows "card A = card B" + using assms by (blast intro: card_inj order_antisym) + + +subsection \Extensional Function Spaces\ + +definition PiE :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" + where "PiE S T = Pi S T \ extensional S" + +abbreviation "Pi\<^sub>E A B \ PiE A B" + +syntax + "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\\<^sub>E _\_./ _)" 10) +translations + "\\<^sub>E x\A. B" \ "CONST Pi\<^sub>E A (\x. B)" + +abbreviation extensional_funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\\<^sub>E" 60) + where "A \\<^sub>E B \ (\\<^sub>E i\A. B)" + +lemma extensional_funcset_def: "extensional_funcset S T = (S \ T) \ extensional S" + by (simp add: PiE_def) + +lemma PiE_empty_domain[simp]: "Pi\<^sub>E {} T = {\x. undefined}" + unfolding PiE_def by simp + +lemma PiE_UNIV_domain: "Pi\<^sub>E UNIV T = Pi UNIV T" + unfolding PiE_def by simp + +lemma PiE_empty_range[simp]: "i \ I \ F i = {} \ (\\<^sub>E i\I. F i) = {}" + unfolding PiE_def by auto + +lemma PiE_eq_empty_iff: "Pi\<^sub>E I F = {} \ (\i\I. F i = {})" +proof + assume "Pi\<^sub>E I F = {}" + show "\i\I. F i = {}" + proof (rule ccontr) + assume "\ ?thesis" + then have "\i. \y. (i \ I \ y \ F i) \ (i \ I \ y = undefined)" + by auto + from choice[OF this] + obtain f where " \x. (x \ I \ f x \ F x) \ (x \ I \ f x = undefined)" .. + then have "f \ Pi\<^sub>E I F" + by (auto simp: extensional_def PiE_def) + with \Pi\<^sub>E I F = {}\ show False + by auto + qed +qed (auto simp: PiE_def) + +lemma PiE_arb: "f \ Pi\<^sub>E S T \ x \ S \ f x = undefined" + unfolding PiE_def by auto (auto dest!: extensional_arb) + +lemma PiE_mem: "f \ Pi\<^sub>E S T \ x \ S \ f x \ T x" + unfolding PiE_def by auto + +lemma PiE_fun_upd: "y \ T x \ f \ Pi\<^sub>E S T \ f(x := y) \ Pi\<^sub>E (insert x S) T" + unfolding PiE_def extensional_def by auto + +lemma fun_upd_in_PiE: "x \ S \ f \ Pi\<^sub>E (insert x S) T \ f(x := undefined) \ Pi\<^sub>E S T" + unfolding PiE_def extensional_def by auto + +lemma PiE_insert_eq: "Pi\<^sub>E (insert x S) T = (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" +proof - + { + fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" + then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" + by (auto intro!: image_eqI[where x="(f x, f(x := undefined))"] intro: fun_upd_in_PiE PiE_mem) + } + moreover + { + fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" + then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" + by (auto intro!: image_eqI[where x="(f x, f)"] intro: fun_upd_in_PiE PiE_mem simp: insert_absorb) + } + ultimately show ?thesis + by (auto intro: PiE_fun_upd) +qed + +lemma PiE_Int: "Pi\<^sub>E I A \ Pi\<^sub>E I B = Pi\<^sub>E I (\x. A x \ B x)" + by (auto simp: PiE_def) + +lemma PiE_cong: "(\i. i\I \ A i = B i) \ Pi\<^sub>E I A = Pi\<^sub>E I B" + unfolding PiE_def by (auto simp: Pi_cong) + +lemma PiE_E [elim]: + assumes "f \ Pi\<^sub>E A B" + obtains "x \ A" and "f x \ B x" + | "x \ A" and "f x = undefined" + using assms by (auto simp: Pi_def PiE_def extensional_def) + +lemma PiE_I[intro!]: + "(\x. x \ A \ f x \ B x) \ (\x. x \ A \ f x = undefined) \ f \ Pi\<^sub>E A B" + by (simp add: PiE_def extensional_def) + +lemma PiE_mono: "(\x. x \ A \ B x \ C x) \ Pi\<^sub>E A B \ Pi\<^sub>E A C" + by auto + +lemma PiE_iff: "f \ Pi\<^sub>E I X \ (\i\I. f i \ X i) \ f \ extensional I" + by (simp add: PiE_def Pi_iff) + +lemma PiE_restrict[simp]: "f \ Pi\<^sub>E A B \ restrict f A = f" + by (simp add: extensional_restrict PiE_def) + +lemma restrict_PiE[simp]: "restrict f I \ Pi\<^sub>E I S \ f \ Pi I S" + by (auto simp: PiE_iff) + +lemma PiE_eq_subset: + assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" + and eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" + and "i \ I" + shows "F i \ F' i" +proof + fix x + assume "x \ F i" + with ne have "\j. \y. (j \ I \ y \ F j \ (i = j \ x = y)) \ (j \ I \ y = undefined)" + by auto + from choice[OF this] obtain f + where f: " \j. (j \ I \ f j \ F j \ (i = j \ x = f j)) \ (j \ I \ f j = undefined)" .. + then have "f \ Pi\<^sub>E I F" + by (auto simp: extensional_def PiE_def) + then have "f \ Pi\<^sub>E I F'" + using assms by simp + then show "x \ F' i" + using f \i \ I\ by (auto simp: PiE_def) +qed + +lemma PiE_eq_iff_not_empty: + assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" + shows "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i)" +proof (intro iffI ballI) + fix i + assume eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" + assume i: "i \ I" + show "F i = F' i" + using PiE_eq_subset[of I F F', OF ne eq i] + using PiE_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] + by auto +qed (auto simp: PiE_def) + +lemma PiE_eq_iff: + "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i) \ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" +proof (intro iffI disjCI) + assume eq[simp]: "Pi\<^sub>E I F = Pi\<^sub>E I F'" + assume "\ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" + then have "(\i\I. F i \ {}) \ (\i\I. F' i \ {})" + using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by auto + with PiE_eq_iff_not_empty[of I F F'] show "\i\I. F i = F' i" + by auto +next + assume "(\i\I. F i = F' i) \ (\i\I. F i = {}) \ (\i\I. F' i = {})" + then show "Pi\<^sub>E I F = Pi\<^sub>E I F'" + using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by (auto simp: PiE_def) +qed + +lemma extensional_funcset_fun_upd_restricts_rangeI: + "\y \ S. f x \ f y \ f \ (insert x S) \\<^sub>E T \ f(x := undefined) \ S \\<^sub>E (T - {f x})" + unfolding extensional_funcset_def extensional_def + apply auto + apply (case_tac "x = xa") + apply auto + done + +lemma extensional_funcset_fun_upd_extends_rangeI: + assumes "a \ T" "f \ S \\<^sub>E (T - {a})" + shows "f(x := a) \ insert x S \\<^sub>E T" + using assms unfolding extensional_funcset_def extensional_def by auto + + +subsubsection \Injective Extensional Function Spaces\ + +lemma extensional_funcset_fun_upd_inj_onI: + assumes "f \ S \\<^sub>E (T - {a})" + and "inj_on f S" + shows "inj_on (f(x := a)) S" + using assms + unfolding extensional_funcset_def by (auto intro!: inj_on_fun_updI) + +lemma extensional_funcset_extend_domain_inj_on_eq: + assumes "x \ S" + shows "{f. f \ (insert x S) \\<^sub>E T \ inj_on f (insert x S)} = + (\(y, g). g(x:=y)) ` {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" + using assms + apply (auto del: PiE_I PiE_E) + apply (auto intro: extensional_funcset_fun_upd_inj_onI + extensional_funcset_fun_upd_extends_rangeI del: PiE_I PiE_E) + apply (auto simp add: image_iff inj_on_def) + apply (rule_tac x="xa x" in exI) + apply (auto intro: PiE_mem del: PiE_I PiE_E) + apply (rule_tac x="xa(x := undefined)" in exI) + apply (auto intro!: extensional_funcset_fun_upd_restricts_rangeI) + apply (auto dest!: PiE_mem split: if_split_asm) + done + +lemma extensional_funcset_extend_domain_inj_onI: + assumes "x \ S" + shows "inj_on (\(y, g). g(x := y)) {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" + using assms + apply (auto intro!: inj_onI) + apply (metis fun_upd_same) + apply (metis assms PiE_arb fun_upd_triv fun_upd_upd) + done + + +subsubsection \Cardinality\ + +lemma finite_PiE: "finite S \ (\i. i \ S \ finite (T i)) \ finite (\\<^sub>E i \ S. T i)" + by (induct S arbitrary: T rule: finite_induct) (simp_all add: PiE_insert_eq) + +lemma inj_combinator: "x \ S \ inj_on (\(y, g). g(x := y)) (T x \ Pi\<^sub>E S T)" +proof (safe intro!: inj_onI ext) + fix f y g z + assume "x \ S" + assume fg: "f \ Pi\<^sub>E S T" "g \ Pi\<^sub>E S T" + assume "f(x := y) = g(x := z)" + then have *: "\i. (f(x := y)) i = (g(x := z)) i" + unfolding fun_eq_iff by auto + from this[of x] show "y = z" by simp + fix i from *[of i] \x \ S\ fg show "f i = g i" + by (auto split: if_split_asm simp: PiE_def extensional_def) +qed + +lemma card_PiE: "finite S \ card (\\<^sub>E i \ S. T i) = (\ i\S. card (T i))" +proof (induct rule: finite_induct) + case empty + then show ?case by auto +next + case (insert x S) + then show ?case + by (simp add: PiE_insert_eq inj_combinator card_image card_cartesian_product) +qed + +end diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/HOL.thy Wed May 02 13:49:38 2018 +0200 @@ -1398,6 +1398,9 @@ lemma if_distrib: "f (if c then x else y) = (if c then f x else f y)" by simp +lemma if_distribR: "(if b then f else g) x = (if b then f x else g x)" + by simp + lemma all_if_distrib: "(\x. if x = a then P x else Q x) \ P a \ (\x. x\a \ Q x)" by auto diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Hull.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hull.thy Wed May 02 13:49:38 2018 +0200 @@ -0,0 +1,85 @@ +(* Title: Hull.thy + Author: Amine Chaieb, University of Cambridge + Author: Jose Divasón + Author: Jesús Aransay + Author: Johannes Hölzl, VU Amsterdam +*) + +theory Hull + imports Main +begin + +subsection \A generic notion of the convex, affine, conic hull, or closed "hull".\ + +definition hull :: "('a set \ bool) \ 'a set \ 'a set" (infixl "hull" 75) + where "S hull s = \{t. S t \ s \ t}" + +lemma hull_same: "S s \ S hull s = s" + unfolding hull_def by auto + +lemma hull_in: "(\T. Ball T S \ S (\T)) \ S (S hull s)" + unfolding hull_def Ball_def by auto + +lemma hull_eq: "(\T. Ball T S \ S (\T)) \ (S hull s) = s \ S s" + using hull_same[of S s] hull_in[of S s] by metis + +lemma hull_hull [simp]: "S hull (S hull s) = S hull s" + unfolding hull_def by blast + +lemma hull_subset[intro]: "s \ (S hull s)" + unfolding hull_def by blast + +lemma hull_mono: "s \ t \ (S hull s) \ (S hull t)" + unfolding hull_def by blast + +lemma hull_antimono: "\x. S x \ T x \ (T hull s) \ (S hull s)" + unfolding hull_def by blast + +lemma hull_minimal: "s \ t \ S t \ (S hull s) \ t" + unfolding hull_def by blast + +lemma subset_hull: "S t \ S hull s \ t \ s \ t" + unfolding hull_def by blast + +lemma hull_UNIV [simp]: "S hull UNIV = UNIV" + unfolding hull_def by auto + +lemma hull_unique: "s \ t \ S t \ (\t'. s \ t' \ S t' \ t \ t') \ (S hull s = t)" + unfolding hull_def by auto + +lemma hull_induct: "(\x. x\ S \ P x) \ Q {x. P x} \ \x\ Q hull S. P x" + using hull_minimal[of S "{x. P x}" Q] + by (auto simp add: subset_eq) + +lemma hull_inc: "x \ S \ x \ P hull S" + by (metis hull_subset subset_eq) + +lemma hull_Un_subset: "(S hull s) \ (S hull t) \ (S hull (s \ t))" + unfolding Un_subset_iff by (metis hull_mono Un_upper1 Un_upper2) + +lemma hull_Un: + assumes T: "\T. Ball T S \ S (\T)" + shows "S hull (s \ t) = S hull (S hull s \ S hull t)" + apply (rule equalityI) + apply (meson hull_mono hull_subset sup.mono) + by (metis hull_Un_subset hull_hull hull_mono) + +lemma hull_Un_left: "P hull (S \ T) = P hull (P hull S \ T)" + apply (rule equalityI) + apply (simp add: Un_commute hull_mono hull_subset sup.coboundedI2) + by (metis Un_subset_iff hull_hull hull_mono hull_subset) + +lemma hull_Un_right: "P hull (S \ T) = P hull (S \ P hull T)" + by (metis hull_Un_left sup.commute) + +lemma hull_insert: + "P hull (insert a S) = P hull (insert a (P hull S))" + by (metis hull_Un_right insert_is_Un) + +lemma hull_redundant_eq: "a \ (S hull s) \ S hull (insert a s) = S hull s" + unfolding hull_def by blast + +lemma hull_redundant: "a \ (S hull s) \ S hull (insert a s) = S hull s" + by (metis hull_redundant_eq) + +end \ No newline at end of file diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Library/FuncSet.thy --- a/src/HOL/Library/FuncSet.thy Wed Apr 18 21:12:50 2018 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,568 +0,0 @@ -(* Title: HOL/Library/FuncSet.thy - Author: Florian Kammueller and Lawrence C Paulson, Lukas Bulwahn -*) - -section \Pi and Function Sets\ - -theory FuncSet - imports Main - abbrevs PiE = "Pi\<^sub>E" - and PIE = "\\<^sub>E" -begin - -definition Pi :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" - where "Pi A B = {f. \x. x \ A \ f x \ B x}" - -definition extensional :: "'a set \ ('a \ 'b) set" - where "extensional A = {f. \x. x \ A \ f x = undefined}" - -definition "restrict" :: "('a \ 'b) \ 'a set \ 'a \ 'b" - where "restrict f A = (\x. if x \ A then f x else undefined)" - -abbreviation funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 60) - where "A \ B \ Pi A (\_. B)" - -syntax - "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\ _\_./ _)" 10) - "_lam" :: "pttrn \ 'a set \ ('a \ 'b) \ ('a \ 'b)" ("(3\_\_./ _)" [0,0,3] 3) -translations - "\ x\A. B" \ "CONST Pi A (\x. B)" - "\x\A. f" \ "CONST restrict (\x. f) A" - -definition "compose" :: "'a set \ ('b \ 'c) \ ('a \ 'b) \ ('a \ 'c)" - where "compose A g f = (\x\A. g (f x))" - - -subsection \Basic Properties of @{term Pi}\ - -lemma Pi_I[intro!]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" - by (simp add: Pi_def) - -lemma Pi_I'[simp]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" - by (simp add:Pi_def) - -lemma funcsetI: "(\x. x \ A \ f x \ B) \ f \ A \ B" - by (simp add: Pi_def) - -lemma Pi_mem: "f \ Pi A B \ x \ A \ f x \ B x" - by (simp add: Pi_def) - -lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" - unfolding Pi_def by auto - -lemma PiE [elim]: "f \ Pi A B \ (f x \ B x \ Q) \ (x \ A \ Q) \ Q" - by (auto simp: Pi_def) - -lemma Pi_cong: "(\w. w \ A \ f w = g w) \ f \ Pi A B \ g \ Pi A B" - by (auto simp: Pi_def) - -lemma funcset_id [simp]: "(\x. x) \ A \ A" - by auto - -lemma funcset_mem: "f \ A \ B \ x \ A \ f x \ B" - by (simp add: Pi_def) - -lemma funcset_image: "f \ A \ B \ f ` A \ B" - by auto - -lemma image_subset_iff_funcset: "F ` A \ B \ F \ A \ B" - by auto - -lemma Pi_eq_empty[simp]: "(\ x \ A. B x) = {} \ (\x\A. B x = {})" - apply (simp add: Pi_def) - apply auto - txt \Converse direction requires Axiom of Choice to exhibit a function - picking an element from each non-empty @{term "B x"}\ - apply (drule_tac x = "\u. SOME y. y \ B u" in spec) - apply auto - apply (cut_tac P = "\y. y \ B x" in some_eq_ex) - apply auto - done - -lemma Pi_empty [simp]: "Pi {} B = UNIV" - by (simp add: Pi_def) - -lemma Pi_Int: "Pi I E \ Pi I F = (\ i\I. E i \ F i)" - by auto - -lemma Pi_UN: - fixes A :: "nat \ 'i \ 'a set" - assumes "finite I" - and mono: "\i n m. i \ I \ n \ m \ A n i \ A m i" - shows "(\n. Pi I (A n)) = (\ i\I. \n. A n i)" -proof (intro set_eqI iffI) - fix f - assume "f \ (\ i\I. \n. A n i)" - then have "\i\I. \n. f i \ A n i" - by auto - from bchoice[OF this] obtain n where n: "f i \ A (n i) i" if "i \ I" for i - by auto - obtain k where k: "n i \ k" if "i \ I" for i - using \finite I\ finite_nat_set_iff_bounded_le[of "n`I"] by auto - have "f \ Pi I (A k)" - proof (intro Pi_I) - fix i - assume "i \ I" - from mono[OF this, of "n i" k] k[OF this] n[OF this] - show "f i \ A k i" by auto - qed - then show "f \ (\n. Pi I (A n))" - by auto -qed auto - -lemma Pi_UNIV [simp]: "A \ UNIV = UNIV" - by (simp add: Pi_def) - -text \Covariance of Pi-sets in their second argument\ -lemma Pi_mono: "(\x. x \ A \ B x \ C x) \ Pi A B \ Pi A C" - by auto - -text \Contravariance of Pi-sets in their first argument\ -lemma Pi_anti_mono: "A' \ A \ Pi A B \ Pi A' B" - by auto - -lemma prod_final: - assumes 1: "fst \ f \ Pi A B" - and 2: "snd \ f \ Pi A C" - shows "f \ (\ z \ A. B z \ C z)" -proof (rule Pi_I) - fix z - assume z: "z \ A" - have "f z = (fst (f z), snd (f z))" - by simp - also have "\ \ B z \ C z" - by (metis SigmaI PiE o_apply 1 2 z) - finally show "f z \ B z \ C z" . -qed - -lemma Pi_split_domain[simp]: "x \ Pi (I \ J) X \ x \ Pi I X \ x \ Pi J X" - by (auto simp: Pi_def) - -lemma Pi_split_insert_domain[simp]: "x \ Pi (insert i I) X \ x \ Pi I X \ x i \ X i" - by (auto simp: Pi_def) - -lemma Pi_cancel_fupd_range[simp]: "i \ I \ x \ Pi I (B(i := b)) \ x \ Pi I B" - by (auto simp: Pi_def) - -lemma Pi_cancel_fupd[simp]: "i \ I \ x(i := a) \ Pi I B \ x \ Pi I B" - by (auto simp: Pi_def) - -lemma Pi_fupd_iff: "i \ I \ f \ Pi I (B(i := A)) \ f \ Pi (I - {i}) B \ f i \ A" - apply auto - apply (drule_tac x=x in Pi_mem) - apply (simp_all split: if_split_asm) - apply (drule_tac x=i in Pi_mem) - apply (auto dest!: Pi_mem) - done - - -subsection \Composition With a Restricted Domain: @{term compose}\ - -lemma funcset_compose: "f \ A \ B \ g \ B \ C \ compose A g f \ A \ C" - by (simp add: Pi_def compose_def restrict_def) - -lemma compose_assoc: - assumes "f \ A \ B" - and "g \ B \ C" - and "h \ C \ D" - shows "compose A h (compose A g f) = compose A (compose B h g) f" - using assms by (simp add: fun_eq_iff Pi_def compose_def restrict_def) - -lemma compose_eq: "x \ A \ compose A g f x = g (f x)" - by (simp add: compose_def restrict_def) - -lemma surj_compose: "f ` A = B \ g ` B = C \ compose A g f ` A = C" - by (auto simp add: image_def compose_eq) - - -subsection \Bounded Abstraction: @{term restrict}\ - -lemma restrict_cong: "I = J \ (\i. i \ J =simp=> f i = g i) \ restrict f I = restrict g J" - by (auto simp: restrict_def fun_eq_iff simp_implies_def) - -lemma restrict_in_funcset: "(\x. x \ A \ f x \ B) \ (\x\A. f x) \ A \ B" - by (simp add: Pi_def restrict_def) - -lemma restrictI[intro!]: "(\x. x \ A \ f x \ B x) \ (\x\A. f x) \ Pi A B" - by (simp add: Pi_def restrict_def) - -lemma restrict_apply[simp]: "(\y\A. f y) x = (if x \ A then f x else undefined)" - by (simp add: restrict_def) - -lemma restrict_apply': "x \ A \ (\y\A. f y) x = f x" - by simp - -lemma restrict_ext: "(\x. x \ A \ f x = g x) \ (\x\A. f x) = (\x\A. g x)" - by (simp add: fun_eq_iff Pi_def restrict_def) - -lemma restrict_UNIV: "restrict f UNIV = f" - by (simp add: restrict_def) - -lemma inj_on_restrict_eq [simp]: "inj_on (restrict f A) A = inj_on f A" - by (simp add: inj_on_def restrict_def) - -lemma Id_compose: "f \ A \ B \ f \ extensional A \ compose A (\y\B. y) f = f" - by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) - -lemma compose_Id: "g \ A \ B \ g \ extensional A \ compose A g (\x\A. x) = g" - by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) - -lemma image_restrict_eq [simp]: "(restrict f A) ` A = f ` A" - by (auto simp add: restrict_def) - -lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \ B)" - unfolding restrict_def by (simp add: fun_eq_iff) - -lemma restrict_fupd[simp]: "i \ I \ restrict (f (i := x)) I = restrict f I" - by (auto simp: restrict_def) - -lemma restrict_upd[simp]: "i \ I \ (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" - by (auto simp: fun_eq_iff) - -lemma restrict_Pi_cancel: "restrict x I \ Pi I A \ x \ Pi I A" - by (auto simp: restrict_def Pi_def) - - -subsection \Bijections Between Sets\ - -text \The definition of @{const bij_betw} is in \Fun.thy\, but most of -the theorems belong here, or need at least @{term Hilbert_Choice}.\ - -lemma bij_betwI: - assumes "f \ A \ B" - and "g \ B \ A" - and g_f: "\x. x\A \ g (f x) = x" - and f_g: "\y. y\B \ f (g y) = y" - shows "bij_betw f A B" - unfolding bij_betw_def -proof - show "inj_on f A" - by (metis g_f inj_on_def) - have "f ` A \ B" - using \f \ A \ B\ by auto - moreover - have "B \ f ` A" - by auto (metis Pi_mem \g \ B \ A\ f_g image_iff) - ultimately show "f ` A = B" - by blast -qed - -lemma bij_betw_imp_funcset: "bij_betw f A B \ f \ A \ B" - by (auto simp add: bij_betw_def) - -lemma inj_on_compose: "bij_betw f A B \ inj_on g B \ inj_on (compose A g f) A" - by (auto simp add: bij_betw_def inj_on_def compose_eq) - -lemma bij_betw_compose: "bij_betw f A B \ bij_betw g B C \ bij_betw (compose A g f) A C" - apply (simp add: bij_betw_def compose_eq inj_on_compose) - apply (auto simp add: compose_def image_def) - done - -lemma bij_betw_restrict_eq [simp]: "bij_betw (restrict f A) A B = bij_betw f A B" - by (simp add: bij_betw_def) - - -subsection \Extensionality\ - -lemma extensional_empty[simp]: "extensional {} = {\x. undefined}" - unfolding extensional_def by auto - -lemma extensional_arb: "f \ extensional A \ x \ A \ f x = undefined" - by (simp add: extensional_def) - -lemma restrict_extensional [simp]: "restrict f A \ extensional A" - by (simp add: restrict_def extensional_def) - -lemma compose_extensional [simp]: "compose A f g \ extensional A" - by (simp add: compose_def) - -lemma extensionalityI: - assumes "f \ extensional A" - and "g \ extensional A" - and "\x. x \ A \ f x = g x" - shows "f = g" - using assms by (force simp add: fun_eq_iff extensional_def) - -lemma extensional_restrict: "f \ extensional A \ restrict f A = f" - by (rule extensionalityI[OF restrict_extensional]) auto - -lemma extensional_subset: "f \ extensional A \ A \ B \ f \ extensional B" - unfolding extensional_def by auto - -lemma inv_into_funcset: "f ` A = B \ (\x\B. inv_into A f x) \ B \ A" - by (unfold inv_into_def) (fast intro: someI2) - -lemma compose_inv_into_id: "bij_betw f A B \ compose A (\y\B. inv_into A f y) f = (\x\A. x)" - apply (simp add: bij_betw_def compose_def) - apply (rule restrict_ext, auto) - done - -lemma compose_id_inv_into: "f ` A = B \ compose B f (\y\B. inv_into A f y) = (\x\B. x)" - apply (simp add: compose_def) - apply (rule restrict_ext) - apply (simp add: f_inv_into_f) - done - -lemma extensional_insert[intro, simp]: - assumes "a \ extensional (insert i I)" - shows "a(i := b) \ extensional (insert i I)" - using assms unfolding extensional_def by auto - -lemma extensional_Int[simp]: "extensional I \ extensional I' = extensional (I \ I')" - unfolding extensional_def by auto - -lemma extensional_UNIV[simp]: "extensional UNIV = UNIV" - by (auto simp: extensional_def) - -lemma restrict_extensional_sub[intro]: "A \ B \ restrict f A \ extensional B" - unfolding restrict_def extensional_def by auto - -lemma extensional_insert_undefined[intro, simp]: - "a \ extensional (insert i I) \ a(i := undefined) \ extensional I" - unfolding extensional_def by auto - -lemma extensional_insert_cancel[intro, simp]: - "a \ extensional I \ a \ extensional (insert i I)" - unfolding extensional_def by auto - - -subsection \Cardinality\ - -lemma card_inj: "f \ A \ B \ inj_on f A \ finite B \ card A \ card B" - by (rule card_inj_on_le) auto - -lemma card_bij: - assumes "f \ A \ B" "inj_on f A" - and "g \ B \ A" "inj_on g B" - and "finite A" "finite B" - shows "card A = card B" - using assms by (blast intro: card_inj order_antisym) - - -subsection \Extensional Function Spaces\ - -definition PiE :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" - where "PiE S T = Pi S T \ extensional S" - -abbreviation "Pi\<^sub>E A B \ PiE A B" - -syntax - "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\\<^sub>E _\_./ _)" 10) -translations - "\\<^sub>E x\A. B" \ "CONST Pi\<^sub>E A (\x. B)" - -abbreviation extensional_funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\\<^sub>E" 60) - where "A \\<^sub>E B \ (\\<^sub>E i\A. B)" - -lemma extensional_funcset_def: "extensional_funcset S T = (S \ T) \ extensional S" - by (simp add: PiE_def) - -lemma PiE_empty_domain[simp]: "Pi\<^sub>E {} T = {\x. undefined}" - unfolding PiE_def by simp - -lemma PiE_UNIV_domain: "Pi\<^sub>E UNIV T = Pi UNIV T" - unfolding PiE_def by simp - -lemma PiE_empty_range[simp]: "i \ I \ F i = {} \ (\\<^sub>E i\I. F i) = {}" - unfolding PiE_def by auto - -lemma PiE_eq_empty_iff: "Pi\<^sub>E I F = {} \ (\i\I. F i = {})" -proof - assume "Pi\<^sub>E I F = {}" - show "\i\I. F i = {}" - proof (rule ccontr) - assume "\ ?thesis" - then have "\i. \y. (i \ I \ y \ F i) \ (i \ I \ y = undefined)" - by auto - from choice[OF this] - obtain f where " \x. (x \ I \ f x \ F x) \ (x \ I \ f x = undefined)" .. - then have "f \ Pi\<^sub>E I F" - by (auto simp: extensional_def PiE_def) - with \Pi\<^sub>E I F = {}\ show False - by auto - qed -qed (auto simp: PiE_def) - -lemma PiE_arb: "f \ Pi\<^sub>E S T \ x \ S \ f x = undefined" - unfolding PiE_def by auto (auto dest!: extensional_arb) - -lemma PiE_mem: "f \ Pi\<^sub>E S T \ x \ S \ f x \ T x" - unfolding PiE_def by auto - -lemma PiE_fun_upd: "y \ T x \ f \ Pi\<^sub>E S T \ f(x := y) \ Pi\<^sub>E (insert x S) T" - unfolding PiE_def extensional_def by auto - -lemma fun_upd_in_PiE: "x \ S \ f \ Pi\<^sub>E (insert x S) T \ f(x := undefined) \ Pi\<^sub>E S T" - unfolding PiE_def extensional_def by auto - -lemma PiE_insert_eq: "Pi\<^sub>E (insert x S) T = (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" -proof - - { - fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" - then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" - by (auto intro!: image_eqI[where x="(f x, f(x := undefined))"] intro: fun_upd_in_PiE PiE_mem) - } - moreover - { - fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" - then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" - by (auto intro!: image_eqI[where x="(f x, f)"] intro: fun_upd_in_PiE PiE_mem simp: insert_absorb) - } - ultimately show ?thesis - by (auto intro: PiE_fun_upd) -qed - -lemma PiE_Int: "Pi\<^sub>E I A \ Pi\<^sub>E I B = Pi\<^sub>E I (\x. A x \ B x)" - by (auto simp: PiE_def) - -lemma PiE_cong: "(\i. i\I \ A i = B i) \ Pi\<^sub>E I A = Pi\<^sub>E I B" - unfolding PiE_def by (auto simp: Pi_cong) - -lemma PiE_E [elim]: - assumes "f \ Pi\<^sub>E A B" - obtains "x \ A" and "f x \ B x" - | "x \ A" and "f x = undefined" - using assms by (auto simp: Pi_def PiE_def extensional_def) - -lemma PiE_I[intro!]: - "(\x. x \ A \ f x \ B x) \ (\x. x \ A \ f x = undefined) \ f \ Pi\<^sub>E A B" - by (simp add: PiE_def extensional_def) - -lemma PiE_mono: "(\x. x \ A \ B x \ C x) \ Pi\<^sub>E A B \ Pi\<^sub>E A C" - by auto - -lemma PiE_iff: "f \ Pi\<^sub>E I X \ (\i\I. f i \ X i) \ f \ extensional I" - by (simp add: PiE_def Pi_iff) - -lemma PiE_restrict[simp]: "f \ Pi\<^sub>E A B \ restrict f A = f" - by (simp add: extensional_restrict PiE_def) - -lemma restrict_PiE[simp]: "restrict f I \ Pi\<^sub>E I S \ f \ Pi I S" - by (auto simp: PiE_iff) - -lemma PiE_eq_subset: - assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" - and eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" - and "i \ I" - shows "F i \ F' i" -proof - fix x - assume "x \ F i" - with ne have "\j. \y. (j \ I \ y \ F j \ (i = j \ x = y)) \ (j \ I \ y = undefined)" - by auto - from choice[OF this] obtain f - where f: " \j. (j \ I \ f j \ F j \ (i = j \ x = f j)) \ (j \ I \ f j = undefined)" .. - then have "f \ Pi\<^sub>E I F" - by (auto simp: extensional_def PiE_def) - then have "f \ Pi\<^sub>E I F'" - using assms by simp - then show "x \ F' i" - using f \i \ I\ by (auto simp: PiE_def) -qed - -lemma PiE_eq_iff_not_empty: - assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" - shows "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i)" -proof (intro iffI ballI) - fix i - assume eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" - assume i: "i \ I" - show "F i = F' i" - using PiE_eq_subset[of I F F', OF ne eq i] - using PiE_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] - by auto -qed (auto simp: PiE_def) - -lemma PiE_eq_iff: - "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i) \ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" -proof (intro iffI disjCI) - assume eq[simp]: "Pi\<^sub>E I F = Pi\<^sub>E I F'" - assume "\ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" - then have "(\i\I. F i \ {}) \ (\i\I. F' i \ {})" - using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by auto - with PiE_eq_iff_not_empty[of I F F'] show "\i\I. F i = F' i" - by auto -next - assume "(\i\I. F i = F' i) \ (\i\I. F i = {}) \ (\i\I. F' i = {})" - then show "Pi\<^sub>E I F = Pi\<^sub>E I F'" - using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by (auto simp: PiE_def) -qed - -lemma extensional_funcset_fun_upd_restricts_rangeI: - "\y \ S. f x \ f y \ f \ (insert x S) \\<^sub>E T \ f(x := undefined) \ S \\<^sub>E (T - {f x})" - unfolding extensional_funcset_def extensional_def - apply auto - apply (case_tac "x = xa") - apply auto - done - -lemma extensional_funcset_fun_upd_extends_rangeI: - assumes "a \ T" "f \ S \\<^sub>E (T - {a})" - shows "f(x := a) \ insert x S \\<^sub>E T" - using assms unfolding extensional_funcset_def extensional_def by auto - - -subsubsection \Injective Extensional Function Spaces\ - -lemma extensional_funcset_fun_upd_inj_onI: - assumes "f \ S \\<^sub>E (T - {a})" - and "inj_on f S" - shows "inj_on (f(x := a)) S" - using assms - unfolding extensional_funcset_def by (auto intro!: inj_on_fun_updI) - -lemma extensional_funcset_extend_domain_inj_on_eq: - assumes "x \ S" - shows "{f. f \ (insert x S) \\<^sub>E T \ inj_on f (insert x S)} = - (\(y, g). g(x:=y)) ` {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" - using assms - apply (auto del: PiE_I PiE_E) - apply (auto intro: extensional_funcset_fun_upd_inj_onI - extensional_funcset_fun_upd_extends_rangeI del: PiE_I PiE_E) - apply (auto simp add: image_iff inj_on_def) - apply (rule_tac x="xa x" in exI) - apply (auto intro: PiE_mem del: PiE_I PiE_E) - apply (rule_tac x="xa(x := undefined)" in exI) - apply (auto intro!: extensional_funcset_fun_upd_restricts_rangeI) - apply (auto dest!: PiE_mem split: if_split_asm) - done - -lemma extensional_funcset_extend_domain_inj_onI: - assumes "x \ S" - shows "inj_on (\(y, g). g(x := y)) {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" - using assms - apply (auto intro!: inj_onI) - apply (metis fun_upd_same) - apply (metis assms PiE_arb fun_upd_triv fun_upd_upd) - done - - -subsubsection \Cardinality\ - -lemma finite_PiE: "finite S \ (\i. i \ S \ finite (T i)) \ finite (\\<^sub>E i \ S. T i)" - by (induct S arbitrary: T rule: finite_induct) (simp_all add: PiE_insert_eq) - -lemma inj_combinator: "x \ S \ inj_on (\(y, g). g(x := y)) (T x \ Pi\<^sub>E S T)" -proof (safe intro!: inj_onI ext) - fix f y g z - assume "x \ S" - assume fg: "f \ Pi\<^sub>E S T" "g \ Pi\<^sub>E S T" - assume "f(x := y) = g(x := z)" - then have *: "\i. (f(x := y)) i = (g(x := z)) i" - unfolding fun_eq_iff by auto - from this[of x] show "y = z" by simp - fix i from *[of i] \x \ S\ fg show "f i = g i" - by (auto split: if_split_asm simp: PiE_def extensional_def) -qed - -lemma card_PiE: "finite S \ card (\\<^sub>E i \ S. T i) = (\ i\S. card (T i))" -proof (induct rule: finite_induct) - case empty - then show ?case by auto -next - case (insert x S) - then show ?case - by (simp add: PiE_insert_eq inj_combinator card_image card_cartesian_product) -qed - -end diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Library/Library.thy Wed May 02 13:49:38 2018 +0200 @@ -28,7 +28,6 @@ Finite_Map Float FSet - FuncSet Function_Division Fun_Lexorder Going_To_Filter diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Metis_Examples/Abstraction.thy --- a/src/HOL/Metis_Examples/Abstraction.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Metis_Examples/Abstraction.thy Wed May 02 13:49:38 2018 +0200 @@ -8,7 +8,7 @@ section \Example Featuring Metis's Support for Lambda-Abstractions\ theory Abstraction -imports "HOL-Library.FuncSet" +imports HOL.FuncSet begin (* For Christoph Benzmüller *) diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Metis_Examples/Tarski.thy --- a/src/HOL/Metis_Examples/Tarski.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Metis_Examples/Tarski.thy Wed May 02 13:49:38 2018 +0200 @@ -8,7 +8,7 @@ section \Metis Example Featuring the Full Theorem of Tarski\ theory Tarski -imports Main "HOL-Library.FuncSet" +imports Main HOL.FuncSet begin declare [[metis_new_skolem]] diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Modules.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Modules.thy Wed May 02 13:49:38 2018 +0200 @@ -0,0 +1,937 @@ +(* Title: Modules.thy + Author: Amine Chaieb, University of Cambridge + Author: Jose Divasón + Author: Jesús Aransay + Author: Johannes Hölzl, VU Amsterdam + Author: Fabian Immler, TUM +*) + +section \Modules\ + +text \Bases of a linear algebra based on modules (i.e. vector spaces of rings). \ + +theory Modules + imports Hull +begin + +subsection \Locale for additive functions\ + +locale additive = + fixes f :: "'a::ab_group_add \ 'b::ab_group_add" + assumes add: "f (x + y) = f x + f y" +begin + +lemma zero: "f 0 = 0" +proof - + have "f 0 = f (0 + 0)" by simp + also have "\ = f 0 + f 0" by (rule add) + finally show "f 0 = 0" by simp +qed + +lemma minus: "f (- x) = - f x" +proof - + have "f (- x) + f x = f (- x + x)" by (rule add [symmetric]) + also have "\ = - f x + f x" by (simp add: zero) + finally show "f (- x) = - f x" by (rule add_right_imp_eq) +qed + +lemma diff: "f (x - y) = f x - f y" + using add [of x "- y"] by (simp add: minus) + +lemma sum: "f (sum g A) = (\x\A. f (g x))" + by (induct A rule: infinite_finite_induct) (simp_all add: zero add) + +end + + +text \Modules form the central spaces in linear algebra. They are a generalization from vector +spaces by replacing the scalar field by a scalar ring.\ +locale module = + fixes scale :: "'a::comm_ring_1 \ 'b::ab_group_add \ 'b" (infixr "*s" 75) + assumes scale_right_distrib [algebra_simps]: "a *s (x + y) = a *s x + a *s y" + and scale_left_distrib [algebra_simps]: "(a + b) *s x = a *s x + b *s x" + and scale_scale [simp]: "a *s (b *s x) = (a * b) *s x" + and scale_one [simp]: "1 *s x = x" +begin + +lemma scale_left_commute: "a *s (b *s x) = b *s (a *s x)" + by (simp add: mult.commute) + +lemma scale_zero_left [simp]: "0 *s x = 0" + and scale_minus_left [simp]: "(- a) *s x = - (a *s x)" + and scale_left_diff_distrib [algebra_simps]: "(a - b) *s x = a *s x - b *s x" + and scale_sum_left: "(sum f A) *s x = (\a\A. (f a) *s x)" +proof - + interpret s: additive "\a. a *s x" + by standard (rule scale_left_distrib) + show "0 *s x = 0" by (rule s.zero) + show "(- a) *s x = - (a *s x)" by (rule s.minus) + show "(a - b) *s x = a *s x - b *s x" by (rule s.diff) + show "(sum f A) *s x = (\a\A. (f a) *s x)" by (rule s.sum) +qed + +lemma scale_zero_right [simp]: "a *s 0 = 0" + and scale_minus_right [simp]: "a *s (- x) = - (a *s x)" + and scale_right_diff_distrib [algebra_simps]: "a *s (x - y) = a *s x - a *s y" + and scale_sum_right: "a *s (sum f A) = (\x\A. a *s (f x))" +proof - + interpret s: additive "\x. a *s x" + by standard (rule scale_right_distrib) + show "a *s 0 = 0" by (rule s.zero) + show "a *s (- x) = - (a *s x)" by (rule s.minus) + show "a *s (x - y) = a *s x - a *s y" by (rule s.diff) + show "a *s (sum f A) = (\x\A. a *s (f x))" by (rule s.sum) +qed + +lemma sum_constant_scale: "(\x\A. y) = scale (of_nat (card A)) y" + by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) + +section \Subspace\ + +definition subspace :: "'b set \ bool" + where "subspace S \ 0 \ S \ (\x\S. \y\S. x + y \ S) \ (\c. \x\S. c *s x \ S)" + +lemma subspaceI: + "0 \ S \ (\x y. x \ S \ y \ S \ x + y \ S) \ (\c x. x \ S \ c *s x \ S) \ subspace S" + by (auto simp: subspace_def) + +lemma subspace_UNIV[simp]: "subspace UNIV" + by (simp add: subspace_def) + +lemma subspace_single_0[simp]: "subspace {0}" + by (simp add: subspace_def) + +lemma subspace_0: "subspace S \ 0 \ S" + by (metis subspace_def) + +lemma subspace_add: "subspace S \ x \ S \ y \ S \ x + y \ S" + by (metis subspace_def) + +lemma subspace_scale: "subspace S \ x \ S \ c *s x \ S" + by (metis subspace_def) + +lemma subspace_neg: "subspace S \ x \ S \ - x \ S" + by (metis scale_minus_left scale_one subspace_scale) + +lemma subspace_diff: "subspace S \ x \ S \ y \ S \ x - y \ S" + by (metis diff_conv_add_uminus subspace_add subspace_neg) + +lemma subspace_sum: "subspace A \ (\x. x \ B \ f x \ A) \ sum f B \ A" + by (induct B rule: infinite_finite_induct) (auto simp add: subspace_add subspace_0) + +lemma subspace_Int: "(\i. i \ I \ subspace (s i)) \ subspace (\i\I. s i)" + by (auto simp: subspace_def) + +lemma subspace_Inter: "\s \ f. subspace s \ subspace (\f)" + unfolding subspace_def by auto + +lemma subspace_inter: "subspace A \ subspace B \ subspace (A \ B)" + by (simp add: subspace_def) + + +section \Span: subspace generated by a set\ + +definition span :: "'b set \ 'b set" + where span_explicit: "span b = {(\a\t. r a *s a) | t r. finite t \ t \ b}" + +lemma span_explicit': + "span b = {(\v | f v \ 0. f v *s v) | f. finite {v. f v \ 0} \ (\v. f v \ 0 \ v \ b)}" + unfolding span_explicit +proof safe + fix t r assume "finite t" "t \ b" + then show "\f. (\a\t. r a *s a) = (\v | f v \ 0. f v *s v) \ finite {v. f v \ 0} \ (\v. f v \ 0 \ v \ b)" + by (intro exI[of _ "\v. if v \ t then r v else 0"]) (auto intro!: sum.mono_neutral_cong_right) +next + fix f :: "'b \ 'a" assume "finite {v. f v \ 0}" "(\v. f v \ 0 \ v \ b)" + then show "\t r. (\v | f v \ 0. f v *s v) = (\a\t. r a *s a) \ finite t \ t \ b" + by (intro exI[of _ "{v. f v \ 0}"] exI[of _ f]) auto +qed + +lemma span_alt: + "span B = {(\x | f x \ 0. f x *s x) | f. {x. f x \ 0} \ B \ finite {x. f x \ 0}}" + unfolding span_explicit' by auto + +lemma span_finite: + assumes fS: "finite S" + shows "span S = range (\u. \v\S. u v *s v)" + unfolding span_explicit +proof safe + fix t r assume "t \ S" then show "(\a\t. r a *s a) \ range (\u. \v\S. u v *s v)" + by (intro image_eqI[of _ _ "\a. if a \ t then r a else 0"]) + (auto simp: if_distrib[of "\r. r *s a" for a] sum.If_cases fS Int_absorb1) +next + show "\t r. (\v\S. u v *s v) = (\a\t. r a *s a) \ finite t \ t \ S" for u + by (intro exI[of _ u] exI[of _ S]) (auto intro: fS) +qed + +lemma span_induct_alt[consumes 1]: + assumes x: "x \ span S" + assumes h0: "h 0" and hS: "\c x y. x \ S \ h y \ h (c *s x + y)" + shows "h x" + using x unfolding span_explicit +proof safe + fix t r assume "finite t" "t \ S" then show "h (\a\t. r a *s a)" + by (induction t) (auto intro!: hS h0) +qed + +lemma span_mono: "A \ B \ span A \ span B" + by (auto simp: span_explicit) + +lemma span_base: "a \ S \ a \ span S" + by (auto simp: span_explicit intro!: exI[of _ "{a}"] exI[of _ "\_. 1"]) + +lemma span_superset: "S \ span S" + by (auto simp: span_base) + +lemma span_zero: "0 \ span S" + by (auto simp: span_explicit intro!: exI[of _ "{}"]) + +lemma span_UNIV: "span UNIV = UNIV" + by (auto intro: span_base) + +lemma span_add: "x \ span S \ y \ span S \ x + y \ span S" + unfolding span_explicit +proof safe + fix tx ty rx ry assume *: "finite tx" "finite ty" "tx \ S" "ty \ S" + have [simp]: "(tx \ ty) \ tx = tx" "(tx \ ty) \ ty = ty" + by auto + show "\t r. (\a\tx. rx a *s a) + (\a\ty. ry a *s a) = (\a\t. r a *s a) \ finite t \ t \ S" + apply (intro exI[of _ "tx \ ty"]) + apply (intro exI[of _ "\a. (if a \ tx then rx a else 0) + (if a \ ty then ry a else 0)"]) + apply (auto simp: * scale_left_distrib sum.distrib if_distrib[of "\r. r *s a" for a] sum.If_cases) + done +qed + +lemma span_scale: "x \ span S \ c *s x \ span S" + unfolding span_explicit +proof safe + fix t r assume *: "finite t" "t \ S" + show "\t' r'. c *s (\a\t. r a *s a) = (\a\t'. r' a *s a) \ finite t' \ t' \ S" + by (intro exI[of _ t] exI[of _ "\a. c * r a"]) (auto simp: * scale_sum_right) +qed + +lemma subspace_span: "subspace (span S)" + by (auto simp: subspace_def span_zero span_add span_scale) + +lemma span_neg: "x \ span S \ - x \ span S" + by (metis subspace_neg subspace_span) + +lemma span_diff: "x \ span S \ y \ span S \ x - y \ span S" + by (metis subspace_span subspace_diff) + +lemma span_sum: "(\x. x \ A \ f x \ span S) \ sum f A \ span S" + by (rule subspace_sum, rule subspace_span) + +lemma span_minimal: "S \ T \ subspace T \ span S \ T" + by (auto simp: span_explicit intro!: subspace_sum subspace_scale) + +lemma span_def: "span S = subspace hull S" + by (intro hull_unique[symmetric] span_superset subspace_span span_minimal) + +lemma span_unique: + "S \ T \ subspace T \ (\T'. S \ T' \ subspace T' \ T \ T') \ span S = T" + unfolding span_def by (rule hull_unique) + +lemma span_subspace_induct[consumes 2]: + assumes x: "x \ span S" + and P: "subspace P" + and SP: "\x. x \ S \ x \ P" + shows "x \ P" +proof - + from SP have SP': "S \ P" + by (simp add: subset_eq) + from x hull_minimal[where S=subspace, OF SP' P, unfolded span_def[symmetric]] + show "x \ P" + by (metis subset_eq) +qed + +lemma (in module) span_induct[consumes 1]: + assumes x: "x \ span S" + and P: "subspace (Collect P)" + and SP: "\x. x \ S \ P x" + shows "P x" + using P SP span_subspace_induct x by fastforce + +lemma (in module) span_induct': + "\x \ S. P x \ subspace {x. P x} \ \x \ span S. P x" + unfolding span_def by (rule hull_induct) auto + +lemma span_empty[simp]: "span {} = {0}" + by (rule span_unique) (auto simp add: subspace_def) + +lemma span_subspace: "A \ B \ B \ span A \ subspace B \ span A = B" + by (metis order_antisym span_def hull_minimal) + +lemma span_span: "span (span A) = span A" + unfolding span_def hull_hull .. + +(* TODO: proof generally for subspace: *) +lemma span_add_eq: assumes x: "x \ span S" shows "x + y \ span S \ y \ span S" +proof + assume *: "x + y \ span S" + have "(x + y) - x \ span S" using * x by (rule span_diff) + then show "y \ span S" by simp +qed (intro span_add x) + +lemma span_add_eq2: assumes y: "y \ span S" shows "x + y \ span S \ x \ span S" + using span_add_eq[of y S x] y by (auto simp: ac_simps) + +lemma span_singleton: "span {x} = range (\k. k *s x)" + by (auto simp: span_finite) + +lemma span_Un: "span (S \ T) = {x + y | x y. x \ span S \ y \ span T}" +proof safe + fix x assume "x \ span (S \ T)" + then obtain t r where t: "finite t" "t \ S \ T" and x: "x = (\a\t. r a *s a)" + by (auto simp: span_explicit) + moreover have "t \ S \ (t - S) = t" by auto + ultimately show "\xa y. x = xa + y \ xa \ span S \ y \ span T" + unfolding x + apply (rule_tac exI[of _ "\a\t \ S. r a *s a"]) + apply (rule_tac exI[of _ "\a\t - S. r a *s a"]) + apply (subst sum.union_inter_neutral[symmetric]) + apply (auto intro!: span_sum span_scale intro: span_base) + done +next + fix x y assume"x \ span S" "y \ span T" then show "x + y \ span (S \ T)" + using span_mono[of S "S \ T"] span_mono[of T "S \ T"] + by (auto intro!: span_add) +qed + +lemma span_insert: "span (insert a S) = {x. \k. (x - k *s a) \ span S}" +proof - + have "span ({a} \ S) = {x. \k. (x - k *s a) \ span S}" + unfolding span_Un span_singleton + apply (auto simp add: set_eq_iff) + subgoal for y k by (auto intro!: exI[of _ "k"]) + subgoal for y k by (rule exI[of _ "k *s a"], rule exI[of _ "y - k *s a"]) auto + done + then show ?thesis by simp +qed + +lemma span_breakdown: + assumes bS: "b \ S" + and aS: "a \ span S" + shows "\k. a - k *s b \ span (S - {b})" + using assms span_insert [of b "S - {b}"] + by (simp add: insert_absorb) + +lemma span_breakdown_eq: "x \ span (insert a S) \ (\k. x - k *s a \ span S)" + by (simp add: span_insert) + +lemmas span_clauses = span_base span_zero span_add span_scale + +lemma span_eq_iff[simp]: "span s = s \ subspace s" + unfolding span_def by (rule hull_eq) (rule subspace_Inter) + +lemma span_eq: "span S = span T \ S \ span T \ T \ span S" + by (metis span_minimal span_subspace span_superset subspace_span) + +lemma eq_span_insert_eq: + assumes "(x - y) \ span S" + shows "span(insert x S) = span(insert y S)" +proof - + have *: "span(insert x S) \ span(insert y S)" if "(x - y) \ span S" for x y + proof - + have 1: "(r *s x - r *s y) \ span S" for r + by (metis scale_right_diff_distrib span_scale that) + have 2: "(z - k *s y) - k *s (x - y) = z - k *s x" for z k + by (simp add: scale_right_diff_distrib) + show ?thesis + apply (clarsimp simp add: span_breakdown_eq) + by (metis 1 2 diff_add_cancel scale_right_diff_distrib span_add_eq) + qed + show ?thesis + apply (intro subset_antisym * assms) + using assms subspace_neg subspace_span minus_diff_eq by force +qed + + +section \Dependent and independent sets\ + +definition dependent :: "'b set \ bool" + where dependent_explicit: "dependent s \ (\t u. finite t \ t \ s \ (\v\t. u v *s v) = 0 \ (\v\t. u v \ 0))" + +abbreviation "independent s \ \ dependent s" + +lemma dependent_mono: "dependent B \ B \ A \ dependent A" + by (auto simp add: dependent_explicit) + +lemma independent_mono: "independent A \ B \ A \ independent B" + by (auto intro: dependent_mono) + +lemma dependent_zero: "0 \ A \ dependent A" + by (auto simp: dependent_explicit intro!: exI[of _ "\i. 1"] exI[of _ "{0}"]) + +lemma independent_empty[intro]: "independent {}" + by (simp add: dependent_explicit) + +lemma independent_explicit_module: + "independent s \ (\t u v. finite t \ t \ s \ (\v\t. u v *s v) = 0 \ v \ t \ u v = 0)" + unfolding dependent_explicit by auto + +lemma independentD: "independent s \ finite t \ t \ s \ (\v\t. u v *s v) = 0 \ v \ t \ u v = 0" + by (simp add: independent_explicit_module) + +lemma independent_Union_directed: + assumes directed: "\c d. c \ C \ d \ C \ c \ d \ d \ c" + assumes indep: "\c. c \ C \ independent c" + shows "independent (\C)" +proof + assume "dependent (\C)" + then obtain u v S where S: "finite S" "S \ \C" "v \ S" "u v \ 0" "(\v\S. u v *s v) = 0" + by (auto simp: dependent_explicit) + + have "S \ {}" + using \v \ S\ by auto + have "\c\C. S \ c" + using \finite S\ \S \ {}\ \S \ \C\ + proof (induction rule: finite_ne_induct) + case (insert i I) + then obtain c d where cd: "c \ C" "d \ C" and iI: "I \ c" "i \ d" + by blast + from directed[OF cd] cd have "c \ d \ C" + by (auto simp: sup.absorb1 sup.absorb2) + with iI show ?case + by (intro bexI[of _ "c \ d"]) auto + qed auto + then obtain c where "c \ C" "S \ c" + by auto + have "dependent c" + unfolding dependent_explicit + by (intro exI[of _ S] exI[of _ u] bexI[of _ v] conjI) fact+ + with indep[OF \c \ C\] show False + by auto +qed + +lemma dependent_finite: + assumes "finite S" + shows "dependent S \ (\u. (\v \ S. u v \ 0) \ (\v\S. u v *s v) = 0)" + (is "?lhs = ?rhs") +proof + assume ?lhs + then obtain T u v + where "finite T" "T \ S" "v\T" "u v \ 0" "(\v\T. u v *s v) = 0" + by (force simp: dependent_explicit) + with assms show ?rhs + apply (rule_tac x="\v. if v \ T then u v else 0" in exI) + apply (auto simp: sum.mono_neutral_right) + done +next + assume ?rhs with assms show ?lhs + by (fastforce simp add: dependent_explicit) +qed + +lemma dependent_alt: + "dependent B \ + (\X. finite {x. X x \ 0} \ {x. X x \ 0} \ B \ (\x|X x \ 0. X x *s x) = 0 \ (\x. X x \ 0))" + unfolding dependent_explicit + apply safe + subgoal for S u v + apply (intro exI[of _ "\x. if x \ S then u x else 0"]) + apply (subst sum.mono_neutral_cong_left[where T=S]) + apply (auto intro!: sum.mono_neutral_cong_right cong: rev_conj_cong) + done + apply auto + done + +lemma independent_alt: + "independent B \ + (\X. finite {x. X x \ 0} \ {x. X x \ 0} \ B \ (\x|X x \ 0. X x *s x) = 0 \ (\x. X x = 0))" + unfolding dependent_alt by auto + +lemma independentD_alt: + "independent B \ finite {x. X x \ 0} \ {x. X x \ 0} \ B \ (\x|X x \ 0. X x *s x) = 0 \ X x = 0" + unfolding independent_alt by blast + +lemma independentD_unique: + assumes B: "independent B" + and X: "finite {x. X x \ 0}" "{x. X x \ 0} \ B" + and Y: "finite {x. Y x \ 0}" "{x. Y x \ 0} \ B" + and "(\x | X x \ 0. X x *s x) = (\x| Y x \ 0. Y x *s x)" + shows "X = Y" +proof - + have "X x - Y x = 0" for x + using B + proof (rule independentD_alt) + have "{x. X x - Y x \ 0} \ {x. X x \ 0} \ {x. Y x \ 0}" + by auto + then show "finite {x. X x - Y x \ 0}" "{x. X x - Y x \ 0} \ B" + using X Y by (auto dest: finite_subset) + then have "(\x | X x - Y x \ 0. (X x - Y x) *s x) = (\v\{S. X S \ 0} \ {S. Y S \ 0}. (X v - Y v) *s v)" + using X Y by (intro sum.mono_neutral_cong_left) auto + also have "\ = (\v\{S. X S \ 0} \ {S. Y S \ 0}. X v *s v) - (\v\{S. X S \ 0} \ {S. Y S \ 0}. Y v *s v)" + by (simp add: scale_left_diff_distrib sum_subtractf assms) + also have "(\v\{S. X S \ 0} \ {S. Y S \ 0}. X v *s v) = (\v\{S. X S \ 0}. X v *s v)" + using X Y by (intro sum.mono_neutral_cong_right) auto + also have "(\v\{S. X S \ 0} \ {S. Y S \ 0}. Y v *s v) = (\v\{S. Y S \ 0}. Y v *s v)" + using X Y by (intro sum.mono_neutral_cong_right) auto + finally show "(\x | X x - Y x \ 0. (X x - Y x) *s x) = 0" + using assms by simp + qed + then show ?thesis + by auto +qed + + +section \Representation of a vector on a specific basis\ + +definition representation :: "'b set \ 'b \ 'b \ 'a" + where "representation basis v = + (if independent basis \ v \ span basis then + SOME f. (\v. f v \ 0 \ v \ basis) \ finite {v. f v \ 0} \ (\v\{v. f v \ 0}. f v *s v) = v + else (\b. 0))" + +lemma unique_representation: + assumes basis: "independent basis" + and in_basis: "\v. f v \ 0 \ v \ basis" "\v. g v \ 0 \ v \ basis" + and [simp]: "finite {v. f v \ 0}" "finite {v. g v \ 0}" + and eq: "(\v\{v. f v \ 0}. f v *s v) = (\v\{v. g v \ 0}. g v *s v)" + shows "f = g" +proof (rule ext, rule ccontr) + fix v assume ne: "f v \ g v" + have "dependent basis" + unfolding dependent_explicit + proof (intro exI conjI) + have *: "{v. f v - g v \ 0} \ {v. f v \ 0} \ {v. g v \ 0}" + by auto + show "finite {v. f v - g v \ 0}" + by (rule finite_subset[OF *]) simp + show "\v\{v. f v - g v \ 0}. f v - g v \ 0" + by (rule bexI[of _ v]) (auto simp: ne) + have "(\v | f v - g v \ 0. (f v - g v) *s v) = + (\v\{v. f v \ 0} \ {v. g v \ 0}. (f v - g v) *s v)" + by (intro sum.mono_neutral_cong_left *) auto + also have "... = + (\v\{v. f v \ 0} \ {v. g v \ 0}. f v *s v) - (\v\{v. f v \ 0} \ {v. g v \ 0}. g v *s v)" + by (simp add: algebra_simps sum_subtractf) + also have "... = (\v | f v \ 0. f v *s v) - (\v | g v \ 0. g v *s v)" + by (intro arg_cong2[where f= "(-)"] sum.mono_neutral_cong_right) auto + finally show "(\v | f v - g v \ 0. (f v - g v) *s v) = 0" + by (simp add: eq) + show "{v. f v - g v \ 0} \ basis" + using in_basis * by auto + qed + with basis show False by auto +qed + +lemma + shows representation_ne_zero: "\b. representation basis v b \ 0 \ b \ basis" + and finite_representation: "finite {b. representation basis v b \ 0}" + and sum_nonzero_representation_eq: + "independent basis \ v \ span basis \ (\b | representation basis v b \ 0. representation basis v b *s b) = v" +proof - + { assume basis: "independent basis" and v: "v \ span basis" + define p where "p f \ + (\v. f v \ 0 \ v \ basis) \ finite {v. f v \ 0} \ (\v\{v. f v \ 0}. f v *s v) = v" for f + obtain t r where *: "finite t" "t \ basis" "(\b\t. r b *s b) = v" + using \v \ span basis\ by (auto simp: span_explicit) + define f where "f b = (if b \ t then r b else 0)" for b + have "p f" + using * by (auto simp: p_def f_def intro!: sum.mono_neutral_cong_left) + have *: "representation basis v = Eps p" by (simp add: p_def[abs_def] representation_def basis v) + from someI[of p f, OF \p f\] have "p (representation basis v)" + unfolding * . } + note * = this + + show "representation basis v b \ 0 \ b \ basis" for b + using * by (cases "independent basis \ v \ span basis") (auto simp: representation_def) + + show "finite {b. representation basis v b \ 0}" + using * by (cases "independent basis \ v \ span basis") (auto simp: representation_def) + + show "independent basis \ v \ span basis \ (\b | representation basis v b \ 0. representation basis v b *s b) = v" + using * by auto +qed + +lemma sum_representation_eq: + "(\b\B. representation basis v b *s b) = v" + if "independent basis" "v \ span basis" "finite B" "basis \ B" +proof - + have "(\b\B. representation basis v b *s b) = + (\b | representation basis v b \ 0. representation basis v b *s b)" + apply (rule sum.mono_neutral_cong) + apply (rule finite_representation) + apply fact + subgoal for b + using that representation_ne_zero[of basis v b] + by auto + subgoal by auto + subgoal by simp + done + also have "\ = v" + by (rule sum_nonzero_representation_eq; fact) + finally show ?thesis . +qed + +lemma representation_eqI: + assumes basis: "independent basis" and b: "v \ span basis" + and ne_zero: "\b. f b \ 0 \ b \ basis" + and finite: "finite {b. f b \ 0}" + and eq: "(\b | f b \ 0. f b *s b) = v" + shows "representation basis v = f" + by (rule unique_representation[OF basis]) + (auto simp: representation_ne_zero finite_representation + sum_nonzero_representation_eq[OF basis b] ne_zero finite eq) + +lemma representation_basis: + assumes basis: "independent basis" and b: "b \ basis" + shows "representation basis b = (\v. if v = b then 1 else 0)" +proof (rule unique_representation[OF basis]) + show "representation basis b v \ 0 \ v \ basis" for v + using representation_ne_zero . + show "finite {v. representation basis b v \ 0}" + using finite_representation . + show "(if v = b then 1 else 0) \ 0 \ v \ basis" for v + by (cases "v = b") (auto simp: b) + have *: "{v. (if v = b then 1 else 0 :: 'a) \ 0} = {b}" + by auto + show "finite {v. (if v = b then 1 else 0) \ 0}" unfolding * by auto + show "(\v | representation basis b v \ 0. representation basis b v *s v) = + (\v | (if v = b then 1 else 0::'a) \ 0. (if v = b then 1 else 0) *s v)" + unfolding * sum_nonzero_representation_eq[OF basis span_base[OF b]] by auto +qed + +lemma representation_zero: "representation basis 0 = (\b. 0)" +proof cases + assume basis: "independent basis" show ?thesis + by (rule representation_eqI[OF basis span_zero]) auto +qed (simp add: representation_def) + +lemma representation_diff: + assumes basis: "independent basis" and v: "v \ span basis" and u: "u \ span basis" + shows "representation basis (u - v) = (\b. representation basis u b - representation basis v b)" +proof (rule representation_eqI[OF basis span_diff[OF u v]]) + let ?R = "representation basis" + note finite_representation[simp] u[simp] v[simp] + have *: "{b. ?R u b - ?R v b \ 0} \ {b. ?R u b \ 0} \ {b. ?R v b \ 0}" + by auto + then show "?R u b - ?R v b \ 0 \ b \ basis" for b + by (auto dest: representation_ne_zero) + show "finite {b. ?R u b - ?R v b \ 0}" + by (intro finite_subset[OF *]) simp_all + have "(\b | ?R u b - ?R v b \ 0. (?R u b - ?R v b) *s b) = + (\b\{b. ?R u b \ 0} \ {b. ?R v b \ 0}. (?R u b - ?R v b) *s b)" + by (intro sum.mono_neutral_cong_left *) auto + also have "... = + (\b\{b. ?R u b \ 0} \ {b. ?R v b \ 0}. ?R u b *s b) - (\b\{b. ?R u b \ 0} \ {b. ?R v b \ 0}. ?R v b *s b)" + by (simp add: algebra_simps sum_subtractf) + also have "... = (\b | ?R u b \ 0. ?R u b *s b) - (\b | ?R v b \ 0. ?R v b *s b)" + by (intro arg_cong2[where f= "(-)"] sum.mono_neutral_cong_right) auto + finally show "(\b | ?R u b - ?R v b \ 0. (?R u b - ?R v b) *s b) = u - v" + by (simp add: sum_nonzero_representation_eq[OF basis]) +qed + +lemma representation_neg: + "independent basis \ v \ span basis \ representation basis (- v) = (\b. - representation basis v b)" + using representation_diff[of basis v 0] by (simp add: representation_zero span_zero) + +lemma representation_add: + "independent basis \ v \ span basis \ u \ span basis \ + representation basis (u + v) = (\b. representation basis u b + representation basis v b)" + using representation_diff[of basis "-v" u] by (simp add: representation_neg representation_diff span_neg) + +lemma representation_sum: + "independent basis \ (\i. i \ I \ v i \ span basis) \ + representation basis (sum v I) = (\b. \i\I. representation basis (v i) b)" + by (induction I rule: infinite_finite_induct) + (auto simp: representation_zero representation_add span_sum) + +lemma representation_scale: + assumes basis: "independent basis" and v: "v \ span basis" + shows "representation basis (r *s v) = (\b. r * representation basis v b)" +proof (rule representation_eqI[OF basis span_scale[OF v]]) + let ?R = "representation basis" + note finite_representation[simp] v[simp] + have *: "{b. r * ?R v b \ 0} \ {b. ?R v b \ 0}" + by auto + then show "r * representation basis v b \ 0 \ b \ basis" for b + using representation_ne_zero by auto + show "finite {b. r * ?R v b \ 0}" + by (intro finite_subset[OF *]) simp_all + have "(\b | r * ?R v b \ 0. (r * ?R v b) *s b) = (\b\{b. ?R v b \ 0}. (r * ?R v b) *s b)" + by (intro sum.mono_neutral_cong_left *) auto + also have "... = r *s (\b | ?R v b \ 0. ?R v b *s b)" + by (simp add: scale_scale[symmetric] scale_sum_right del: scale_scale) + finally show "(\b | r * ?R v b \ 0. (r * ?R v b) *s b) = r *s v" + by (simp add: sum_nonzero_representation_eq[OF basis]) +qed + +lemma representation_extend: + assumes basis: "independent basis" and v: "v \ span basis'" and basis': "basis' \ basis" + shows "representation basis v = representation basis' v" +proof (rule representation_eqI[OF basis]) + show v': "v \ span basis" using span_mono[OF basis'] v by auto + have *: "independent basis'" using basis' basis by (auto intro: dependent_mono) + show "representation basis' v b \ 0 \ b \ basis" for b + using representation_ne_zero basis' by auto + show "finite {b. representation basis' v b \ 0}" + using finite_representation . + show "(\b | representation basis' v b \ 0. representation basis' v b *s b) = v" + using sum_nonzero_representation_eq[OF * v] . +qed + +text \The set \B\ is the maximal independent set for \span B\, or \A\ is the minimal spanning set\ +lemma spanning_subset_independent: + assumes BA: "B \ A" + and iA: "independent A" + and AsB: "A \ span B" + shows "A = B" +proof (intro antisym[OF _ BA] subsetI) + have iB: "independent B" using independent_mono [OF iA BA] . + fix v assume "v \ A" + with AsB have "v \ span B" by auto + let ?RB = "representation B v" and ?RA = "representation A v" + have "?RB v = 1" + unfolding representation_extend[OF iA \v \ span B\ BA, symmetric] representation_basis[OF iA \v \ A\] by simp + then show "v \ B" + using representation_ne_zero[of B v v] by auto +qed + +end + +(* We need to introduce more specific modules, where the ring structure gets more and more finer, + i.e. Bezout rings & domains, division rings, fields *) + +text \A linear function is a mapping between two modules over the same ring.\ + +locale module_hom = m1: module s1 + m2: module s2 + for s1 :: "'a::comm_ring_1 \ 'b::ab_group_add \ 'b" (infixr "*a" 75) + and s2 :: "'a::comm_ring_1 \ 'c::ab_group_add \ 'c" (infixr "*b" 75) + + fixes f :: "'b \ 'c" + assumes add: "f (b1 + b2) = f b1 + f b2" + and scale: "f (r *a b) = r *b f b" +begin + +lemma zero[simp]: "f 0 = 0" + using scale[of 0 0] by simp + +lemma neg: "f (- x) = - f x" + using scale [where r="-1"] by (metis add add_eq_0_iff zero) + +lemma diff: "f (x - y) = f x - f y" + by (metis diff_conv_add_uminus add neg) + +lemma sum: "f (sum g S) = (\a\S. f (g a))" +proof (induct S rule: infinite_finite_induct) + case (insert x F) + have "f (sum g (insert x F)) = f (g x + sum g F)" + using insert.hyps by simp + also have "\ = f (g x) + f (sum g F)" + using add by simp + also have "\ = (\a\insert x F. f (g a))" + using insert.hyps by simp + finally show ?case . +qed simp_all + +lemma inj_on_iff_eq_0: + assumes s: "m1.subspace s" + shows "inj_on f s \ (\x\s. f x = 0 \ x = 0)" +proof - + have "inj_on f s \ (\x\s. \y\s. f x - f y = 0 \ x - y = 0)" + by (simp add: inj_on_def) + also have "\ \ (\x\s. \y\s. f (x - y) = 0 \ x - y = 0)" + by (simp add: diff) + also have "\ \ (\x\s. f x = 0 \ x = 0)" (is "?l = ?r")(* TODO: sledgehammer! *) + proof safe + fix x assume ?l assume "x \ s" "f x = 0" with \?l\[rule_format, of x 0] s show "x = 0" + by (auto simp: m1.subspace_0) + next + fix x y assume ?r assume "x \ s" "y \ s" "f (x - y) = 0" + with \?r\[rule_format, of "x - y"] s + show "x - y = 0" + by (auto simp: m1.subspace_diff) + qed + finally show ?thesis + by auto +qed + +lemma inj_iff_eq_0: "inj f = (\x. f x = 0 \ x = 0)" + by (rule inj_on_iff_eq_0[OF m1.subspace_UNIV, unfolded ball_UNIV]) + +lemma subspace_image: assumes S: "m1.subspace S" shows "m2.subspace (f ` S)" + unfolding m2.subspace_def +proof safe + show "0 \ f ` S" + by (rule image_eqI[of _ _ 0]) (auto simp: S m1.subspace_0) + show "x \ S \ y \ S \ f x + f y \ f ` S" for x y + by (rule image_eqI[of _ _ "x + y"]) (auto simp: S m1.subspace_add add) + show "x \ S \ r *b f x \ f ` S" for r x + by (rule image_eqI[of _ _ "r *a x"]) (auto simp: S m1.subspace_scale scale) +qed + +lemma subspace_vimage: "m2.subspace S \ m1.subspace (f -` S)" + by (simp add: vimage_def add scale m1.subspace_def m2.subspace_0 m2.subspace_add m2.subspace_scale) + +lemma subspace_kernel: "m1.subspace {x. f x = 0}" + using subspace_vimage[OF m2.subspace_single_0] by (simp add: vimage_def) + +lemma span_image: "m2.span (f ` S) = f ` (m1.span S)" +proof (rule m2.span_unique) + show "f ` S \ f ` m1.span S" + by (rule image_mono, rule m1.span_superset) + show "m2.subspace (f ` m1.span S)" + using m1.subspace_span by (rule subspace_image) +next + fix T assume "f ` S \ T" and "m2.subspace T" then show "f ` m1.span S \ T" + unfolding image_subset_iff_subset_vimage by (metis subspace_vimage m1.span_minimal) +qed + +lemma dependent_inj_imageD: + assumes d: "m2.dependent (f ` s)" and i: "inj_on f (m1.span s)" + shows "m1.dependent s" +proof - + have [intro]: "inj_on f s" + using \inj_on f (m1.span s)\ m1.span_superset by (rule inj_on_subset) + from d obtain s' r v where *: "finite s'" "s' \ s" "(\v\f ` s'. r v *b v) = 0" "v \ s'" "r (f v) \ 0" + by (auto simp: m2.dependent_explicit subset_image_iff dest!: finite_imageD intro: inj_on_subset) + have "f (\v\s'. r (f v) *a v) = (\v\s'. r (f v) *b f v)" + by (simp add: sum scale) + also have "... = (\v\f ` s'. r v *b v)" + using \s' \ s\ by (subst sum.reindex) (auto dest!: finite_imageD intro: inj_on_subset) + finally have "f (\v\s'. r (f v) *a v) = 0" + by (simp add: *) + with \s' \ s\ have "(\v\s'. r (f v) *a v) = 0" + by (intro inj_onD[OF i] m1.span_zero m1.span_sum m1.span_scale) (auto intro: m1.span_base) + then show "m1.dependent s" + using \finite s'\ \s' \ s\ \v \ s'\ \r (f v) \ 0\ by (force simp add: m1.dependent_explicit) +qed + +lemma eq_0_on_span: + assumes f0: "\x. x \ b \ f x = 0" and x: "x \ m1.span b" shows "f x = 0" + using m1.span_induct[OF x subspace_kernel] f0 by simp + +lemma independent_injective_image: "m1.independent s \ inj_on f (m1.span s) \ m2.independent (f ` s)" + using dependent_inj_imageD[of s] by auto + +lemma inj_on_span_independent_image: + assumes ifB: "m2.independent (f ` B)" and f: "inj_on f B" shows "inj_on f (m1.span B)" + unfolding inj_on_iff_eq_0[OF m1.subspace_span] unfolding m1.span_explicit' +proof safe + fix r assume fr: "finite {v. r v \ 0}" and r: "\v. r v \ 0 \ v \ B" + and eq0: "f (\v | r v \ 0. r v *a v) = 0" + have "0 = (\v | r v \ 0. r v *b f v)" + using eq0 by (simp add: sum scale) + also have "... = (\v\f ` {v. r v \ 0}. r (the_inv_into B f v) *b v)" + using r by (subst sum.reindex) (auto simp: the_inv_into_f_f[OF f] intro!: inj_on_subset[OF f] sum.cong) + finally have "r v \ 0 \ r (the_inv_into B f (f v)) = 0" for v + using fr r ifB[unfolded m2.independent_explicit_module, rule_format, + of "f ` {v. r v \ 0}" "\v. r (the_inv_into B f v)"] + by auto + then have "r v = 0" for v + using the_inv_into_f_f[OF f] r by auto + then show "(\v | r v \ 0. r v *a v) = 0" by auto +qed + +lemma inj_on_span_iff_independent_image: "m2.independent (f ` B) \ inj_on f (m1.span B) \ inj_on f B" + using inj_on_span_independent_image[of B] inj_on_subset[OF _ m1.span_superset, of f B] by auto + +lemma subspace_linear_preimage: "m2.subspace S \ m1.subspace {x. f x \ S}" + by (simp add: add scale m1.subspace_def m2.subspace_def) + +lemma spans_image: "V \ m1.span B \ f ` V \ m2.span (f ` B)" + by (metis image_mono span_image) + +text \Relation between bases and injectivity/surjectivity of map.\ + +lemma spanning_surjective_image: + assumes us: "UNIV \ m1.span S" + and sf: "surj f" + shows "UNIV \ m2.span (f ` S)" +proof - + have "UNIV \ f ` UNIV" + using sf by (auto simp add: surj_def) + also have " \ \ m2.span (f ` S)" + using spans_image[OF us] . + finally show ?thesis . +qed + +lemmas independent_inj_on_image = independent_injective_image + +lemma independent_inj_image: + "m1.independent S \ inj f \ m2.independent (f ` S)" + using independent_inj_on_image[of S] by (auto simp: subset_inj_on) + +end + +lemma module_hom_iff: + "module_hom s1 s2 f \ + module s1 \ module s2 \ + (\x y. f (x + y) = f x + f y) \ (\c x. f (s1 c x) = s2 c (f x))" + by (simp add: module_hom_def module_hom_axioms_def) + +locale module_pair = m1: module s1 + m2: module s2 + for s1 :: "'a :: comm_ring_1 \ 'b \ 'b :: ab_group_add" + and s2 :: "'a :: comm_ring_1 \ 'c \ 'c :: ab_group_add" +begin + +lemma module_hom_zero: "module_hom s1 s2 (\x. 0)" + by (simp add: module_hom_iff m1.module_axioms m2.module_axioms) + +lemma module_hom_add: "module_hom s1 s2 f \ module_hom s1 s2 g \ module_hom s1 s2 (\x. f x + g x)" + by (simp add: module_hom_iff module.scale_right_distrib) + +lemma module_hom_sub: "module_hom s1 s2 f \ module_hom s1 s2 g \ module_hom s1 s2 (\x. f x - g x)" + by (simp add: module_hom_iff module.scale_right_diff_distrib) + +lemma module_hom_neg: "module_hom s1 s2 f \ module_hom s1 s2 (\x. - f x)" + by (simp add: module_hom_iff module.scale_minus_right) + +lemma module_hom_scale: "module_hom s1 s2 f \ module_hom s1 s2 (\x. s2 c (f x))" + by (simp add: module_hom_iff module.scale_scale module.scale_right_distrib ac_simps) + +lemma module_hom_compose_scale: + "module_hom s1 s2 (\x. s2 (f x) (c))" + if "module_hom s1 ( *) f" +proof - + interpret mh: module_hom s1 "( *)" f by fact + show ?thesis + by unfold_locales (simp_all add: mh.add mh.scale m2.scale_left_distrib) +qed + +lemma bij_module_hom_imp_inv_module_hom: "module_hom scale1 scale2 f \ bij f \ + module_hom scale2 scale1 (inv f)" + by (auto simp: module_hom_iff bij_is_surj bij_is_inj surj_f_inv_f + intro!: Hilbert_Choice.inv_f_eq) + +lemma module_hom_sum: "(\i. i \ I \ module_hom s1 s2 (f i)) \ (I = {} \ module s1 \ module s2) \ module_hom s1 s2 (\x. \i\I. f i x)" + apply (induction I rule: infinite_finite_induct) + apply (auto intro!: module_hom_zero module_hom_add) + using m1.module_axioms m2.module_axioms by blast + +lemma module_hom_eq_on_span: "f x = g x" + if "module_hom s1 s2 f" "module_hom s1 s2 g" + and "(\x. x \ B \ f x = g x)" "x \ m1.span B" +proof - + interpret module_hom s1 s2 "\x. f x - g x" + by (rule module_hom_sub that)+ + from eq_0_on_span[OF _ that(4)] that(3) show ?thesis by auto +qed + +end + +context module begin + +lemma module_hom_scale_self[simp]: + "module_hom scale scale (\x. scale c x)" + using module_axioms module_hom_iff scale_left_commute scale_right_distrib by blast + +lemma module_hom_scale_left[simp]: + "module_hom ( *) scale (\r. scale r x)" + by unfold_locales (auto simp: algebra_simps) + +lemma module_hom_id: "module_hom scale scale id" + by (simp add: module_hom_iff module_axioms) + +lemma module_hom_ident: "module_hom scale scale (\x. x)" + by (simp add: module_hom_iff module_axioms) + +lemma module_hom_uminus: "module_hom scale scale uminus" + by (simp add: module_hom_iff module_axioms) + +end + +lemma module_hom_compose: "module_hom s1 s2 f \ module_hom s2 s3 g \ module_hom s1 s3 (g o f)" + by (auto simp: module_hom_iff) + +end diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Number_Theory/Prime_Powers.thy --- a/src/HOL/Number_Theory/Prime_Powers.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Number_Theory/Prime_Powers.thy Wed May 02 13:49:38 2018 +0200 @@ -6,7 +6,7 @@ *) section \Prime powers\ theory Prime_Powers - imports Complex_Main "HOL-Computational_Algebra.Primes" "HOL-Library.FuncSet" + imports Complex_Main "HOL-Computational_Algebra.Primes" begin definition aprimedivisor :: "'a :: normalization_semidom \ 'a" where diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/Real_Vector_Spaces.thy Wed May 02 13:49:38 2018 +0200 @@ -6,123 +6,9 @@ section \Vector Spaces and Algebras over the Reals\ theory Real_Vector_Spaces -imports Real Topological_Spaces -begin - -subsection \Locale for additive functions\ - -locale additive = - fixes f :: "'a::ab_group_add \ 'b::ab_group_add" - assumes add: "f (x + y) = f x + f y" -begin - -lemma zero: "f 0 = 0" -proof - - have "f 0 = f (0 + 0)" by simp - also have "\ = f 0 + f 0" by (rule add) - finally show "f 0 = 0" by simp -qed - -lemma minus: "f (- x) = - f x" -proof - - have "f (- x) + f x = f (- x + x)" by (rule add [symmetric]) - also have "\ = - f x + f x" by (simp add: zero) - finally show "f (- x) = - f x" by (rule add_right_imp_eq) -qed - -lemma diff: "f (x - y) = f x - f y" - using add [of x "- y"] by (simp add: minus) - -lemma sum: "f (sum g A) = (\x\A. f (g x))" - by (induct A rule: infinite_finite_induct) (simp_all add: zero add) - -end - - -subsection \Vector spaces\ - -locale vector_space = - fixes scale :: "'a::field \ 'b::ab_group_add \ 'b" - assumes scale_right_distrib [algebra_simps]: "scale a (x + y) = scale a x + scale a y" - and scale_left_distrib [algebra_simps]: "scale (a + b) x = scale a x + scale b x" - and scale_scale [simp]: "scale a (scale b x) = scale (a * b) x" - and scale_one [simp]: "scale 1 x = x" +imports Real Topological_Spaces Vector_Spaces begin -lemma scale_left_commute: "scale a (scale b x) = scale b (scale a x)" - by (simp add: mult.commute) - -lemma scale_zero_left [simp]: "scale 0 x = 0" - and scale_minus_left [simp]: "scale (- a) x = - (scale a x)" - and scale_left_diff_distrib [algebra_simps]: "scale (a - b) x = scale a x - scale b x" - and scale_sum_left: "scale (sum f A) x = (\a\A. scale (f a) x)" -proof - - interpret s: additive "\a. scale a x" - by standard (rule scale_left_distrib) - show "scale 0 x = 0" by (rule s.zero) - show "scale (- a) x = - (scale a x)" by (rule s.minus) - show "scale (a - b) x = scale a x - scale b x" by (rule s.diff) - show "scale (sum f A) x = (\a\A. scale (f a) x)" by (rule s.sum) -qed - -lemma scale_zero_right [simp]: "scale a 0 = 0" - and scale_minus_right [simp]: "scale a (- x) = - (scale a x)" - and scale_right_diff_distrib [algebra_simps]: "scale a (x - y) = scale a x - scale a y" - and scale_sum_right: "scale a (sum f A) = (\x\A. scale a (f x))" -proof - - interpret s: additive "\x. scale a x" - by standard (rule scale_right_distrib) - show "scale a 0 = 0" by (rule s.zero) - show "scale a (- x) = - (scale a x)" by (rule s.minus) - show "scale a (x - y) = scale a x - scale a y" by (rule s.diff) - show "scale a (sum f A) = (\x\A. scale a (f x))" by (rule s.sum) -qed - -lemma scale_eq_0_iff [simp]: "scale a x = 0 \ a = 0 \ x = 0" -proof (cases "a = 0") - case True - then show ?thesis by simp -next - case False - have "x = 0" if "scale a x = 0" - proof - - from False that have "scale (inverse a) (scale a x) = 0" by simp - with False show ?thesis by simp - qed - then show ?thesis by force -qed - -lemma scale_left_imp_eq: - assumes nonzero: "a \ 0" - and scale: "scale a x = scale a y" - shows "x = y" -proof - - from scale have "scale a (x - y) = 0" - by (simp add: scale_right_diff_distrib) - with nonzero have "x - y = 0" by simp - then show "x = y" by (simp only: right_minus_eq) -qed - -lemma scale_right_imp_eq: - assumes nonzero: "x \ 0" - and scale: "scale a x = scale b x" - shows "a = b" -proof - - from scale have "scale (a - b) x = 0" - by (simp add: scale_left_diff_distrib) - with nonzero have "a - b = 0" by simp - then show "a = b" by (simp only: right_minus_eq) -qed - -lemma scale_cancel_left [simp]: "scale a x = scale a y \ x = y \ a = 0" - by (auto intro: scale_left_imp_eq) - -lemma scale_cancel_right [simp]: "scale a x = scale b x \ a = b \ x = 0" - by (auto intro: scale_right_imp_eq) - -end - - subsection \Real vector spaces\ class scaleR = @@ -140,13 +26,74 @@ and scaleR_scaleR: "scaleR a (scaleR b x) = scaleR (a * b) x" and scaleR_one: "scaleR 1 x = x" -interpretation real_vector: vector_space "scaleR :: real \ 'a \ 'a::real_vector" + +class real_algebra = real_vector + ring + + assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)" + and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (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 ( *\<^sub>R) ( *\<^sub>R) = linear" + and "Vector_Spaces.linear ( * ) ( *\<^sub>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 + apply unfold_locales + apply (rule scaleR_add_right) + apply (rule scaleR_add_left) + apply (rule scaleR_scaleR) + apply (rule scaleR_one) + apply (force simp add: linear_def) + apply (force simp add: linear_def real_scaleR_def[abs_def]) + done + +hide_const (open)\\locale constants\ + real_vector.dependent + real_vector.independent + real_vector.representation + real_vector.subspace + real_vector.span + real_vector.extend_basis + real_vector.dim + +abbreviation "independent x \ \ dependent x" + +global_interpretation real_vector?: vector_space_pair "scaleR::_\_\'a::real_vector" "scaleR::_\_\'b::real_vector" + rewrites "Vector_Spaces.linear ( *\<^sub>R) ( *\<^sub>R) = linear" + and "Vector_Spaces.linear ( * ) ( *\<^sub>R) = linear" + defines construct_raw_def: construct = real_vector.construct apply unfold_locales - apply (rule scaleR_add_right) - apply (rule scaleR_add_left) - apply (rule scaleR_scaleR) - apply (rule scaleR_one) - done + unfolding linear_def real_scaleR_def + by (rule refl)+ + +hide_const (open)\\locale constants\ + real_vector.construct + +lemma linear_compose: "linear f \ linear g \ linear (g o f)" + unfolding linear_def by (rule Vector_Spaces.linear_compose) text \Recover original theorem names\ @@ -172,6 +119,16 @@ lemmas scaleR_left_diff_distrib = scaleR_diff_left lemmas scaleR_right_diff_distrib = scaleR_diff_right +lemmas linear_injective_0 = linear_inj_iff_eq_0 + and linear_injective_on_subspace_0 = linear_inj_on_iff_eq_0 + and linear_cmul = linear_scale + and linear_scaleR = linear_scale_self + and subspace_mul = subspace_scale + and span_linear_image = linear_span_image + and span_0 = span_zero + and span_mul = span_scale + and injective_scaleR = injective_scale + lemma scaleR_minus1_left [simp]: "scaleR (-1) x = - x" for x :: "'a::real_vector" using scaleR_minus_left [of 1 x] by simp @@ -191,26 +148,6 @@ by simp qed -class real_algebra = real_vector + ring + - assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)" - and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (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 - interpretation scaleR_left: additive "(\a. scaleR a x :: 'a::real_vector)" by standard (rule scaleR_left_distrib) @@ -231,9 +168,7 @@ apply (erule (1) nonzero_inverse_scaleR_distrib) done -lemma sum_constant_scaleR: "(\x\A. y) = of_nat (card A) *\<^sub>R y" - for y :: "'a::real_vector" - by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps) +lemmas sum_constant_scaleR = real_vector.sum_constant_scale\\legacy name\ named_theorems vector_add_divide_simps "to simplify sums of scaled vectors" @@ -247,7 +182,7 @@ "(a / z) *\<^sub>R v - w = (if z = 0 then -w else (a *\<^sub>R v - z *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v - b *\<^sub>R w = (if z = 0 then -b *\<^sub>R w else (a *\<^sub>R v - (b * z) *\<^sub>R w) /\<^sub>R z)" for v :: "'a :: real_vector" - by (simp_all add: divide_inverse_commute scaleR_add_right real_vector.scale_right_diff_distrib) + by (simp_all add: divide_inverse_commute scaleR_add_right scaleR_diff_right) lemma eq_vector_fraction_iff [vector_add_divide_simps]: @@ -271,11 +206,11 @@ then have "inverse m *\<^sub>R (m *\<^sub>R x) = inverse m *\<^sub>R (y - c)" by simp then show "x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" using m0 - by (simp add: real_vector.scale_right_diff_distrib) + by (simp add: scaleR_diff_right) next assume ?rhs with m0 show "m *\<^sub>R x + c = y" - by (simp add: real_vector.scale_right_diff_distrib) + by (simp add: scaleR_diff_right) qed lemma real_vector_eq_affinity: "m \ 0 \ y = m *\<^sub>R x + c \ inverse m *\<^sub>R y - (inverse m *\<^sub>R c) = x" @@ -1402,13 +1337,26 @@ subsection \Bounded Linear and Bilinear Operators\ -locale linear = additive f for f :: "'a::real_vector \ 'b::real_vector" + - assumes scaleR: "f (scaleR r x) = scaleR r (f x)" +lemma linearI: "linear f" + if "\b1 b2. f (b1 + b2) = f b1 + f b2" + "\r b. f (r *\<^sub>R b) = r *\<^sub>R f b" + using that + by unfold_locales (auto simp: algebra_simps) -lemma linear_imp_scaleR: - assumes "linear D" - obtains d where "D = (\x. x *\<^sub>R d)" - by (metis assms linear.scaleR mult.commute mult.left_neutral real_scaleR_def) +lemma linear_iff: + "linear f \ (\x y. f (x + y) = f x + f y) \ (\c x. f (c *\<^sub>R x) = c *\<^sub>R f x)" + (is "linear f \ ?rhs") +proof + assume "linear f" + then interpret f: linear f . + show "?rhs" by (simp add: f.add f.scale) +next + assume "?rhs" + then show "linear f" by (intro linearI) auto +qed + +lemmas linear_scaleR_left = linear_scale_left +lemmas linear_imp_scaleR = linear_imp_scale corollary real_linearD: fixes f :: "real \ real" @@ -1416,14 +1364,8 @@ by (rule linear_imp_scaleR [OF assms]) (force simp: scaleR_conv_of_real) lemma linear_times_of_real: "linear (\x. a * of_real x)" - apply (simp add: linear_def Real_Vector_Spaces.additive_def linear_axioms_def) - by (metis distrib_left mult_scaleR_right scaleR_conv_of_real) - -lemma linearI: - assumes "\x y. f (x + y) = f x + f y" - and "\c x. f (c *\<^sub>R x) = c *\<^sub>R f x" - shows "linear f" - by standard (rule assms)+ + by (auto intro!: linearI simp: distrib_left) + (metis mult_scaleR_right scaleR_conv_of_real) locale bounded_linear = linear f for f :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" @@ -1720,10 +1662,6 @@ by (auto intro: exI[of _ "f 1"] bounded_linear_mult_left) qed -lemma bij_linear_imp_inv_linear: "linear f \ bij f \ linear (inv f)" - by (auto simp: linear_def linear_axioms_def additive_def bij_is_surj bij_is_inj surj_f_inv_f - intro!: Hilbert_Choice.inv_f_eq) - instance real_normed_algebra_1 \ perfect_space proof show "\ open {x}" for x :: 'a diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/Vector_Spaces.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Vector_Spaces.thy Wed May 02 13:49:38 2018 +0200 @@ -0,0 +1,1759 @@ +(* Title: Vector_Spaces.thy + Author: Amine Chaieb, University of Cambridge + Author: Jose Divasón + Author: Jesús Aransay + Author: Johannes Hölzl, VU Amsterdam + Author: Fabian Immler, TUM +*) + +section \Vector Spaces\ + +theory Vector_Spaces + imports Modules FuncSet +begin + +lemma isomorphism_expand: + "f \ g = id \ g \ f = id \ (\x. f (g x) = x) \ (\x. g (f x) = x)" + by (simp add: fun_eq_iff o_def id_def) + +lemma left_right_inverse_eq: + assumes fg: "f \ g = id" + and gh: "g \ h = id" + shows "f = h" +proof - + have "f = f \ (g \ h)" + unfolding gh by simp + also have "\ = (f \ g) \ h" + by (simp add: o_assoc) + finally show "f = h" + unfolding fg by simp +qed + +lemma ordLeq3_finite_infinite: + assumes A: "finite A" and B: "infinite B" shows "ordLeq3 (card_of A) (card_of B)" +proof - + have \ordLeq3 (card_of A) (card_of B) \ ordLeq3 (card_of B) (card_of A)\ + by (intro ordLeq_total card_of_Well_order) + moreover have "\ ordLeq3 (card_of B) (card_of A)" + using B A card_of_ordLeq_finite[of B A] by auto + ultimately show ?thesis by auto +qed + +locale vector_space = + fixes scale :: "'a::field \ 'b::ab_group_add \ 'b" (infixr "*s" 75) + assumes vector_space_assms:\\re-stating the assumptions of \module\ instead of extending \module\ + allows us to rewrite in the sublocale.\ + "a *s (x + y) = a *s x + a *s y" + "(a + b) *s x = a *s x + b *s x" + "a *s (b *s x) = (a * b) *s x" + "1 *s x = x" + +lemma module_iff_vector_space: "module s \ vector_space s" + unfolding module_def vector_space_def .. + +locale linear = vs1: vector_space s1 + vs2: vector_space s2 + module_hom s1 s2 f + for s1 :: "'a::field \ 'b::ab_group_add \ 'b" (infixr "*a" 75) + and s2 :: "'a::field \ 'c::ab_group_add \ 'c" (infixr "*b" 75) + and f :: "'b \ 'c" + +lemma module_hom_iff_linear: "module_hom s1 s2 f \ linear s1 s2 f" + unfolding module_hom_def linear_def module_iff_vector_space by auto +lemmas module_hom_eq_linear = module_hom_iff_linear[abs_def, THEN meta_eq_to_obj_eq] +lemmas linear_iff_module_hom = module_hom_iff_linear[symmetric] +lemmas linear_module_homI = module_hom_iff_linear[THEN iffD1] + and module_hom_linearI = module_hom_iff_linear[THEN iffD2] + +context vector_space begin + +sublocale module scale rewrites "module_hom = linear" + by (unfold_locales) (fact vector_space_assms module_hom_eq_linear)+ + +lemmas\\from \module\\ + linear_id = module_hom_id + and linear_ident = module_hom_ident + and linear_scale_self = module_hom_scale_self + and linear_scale_left = module_hom_scale_left + and linear_uminus = module_hom_uminus + +lemma linear_imp_scale: + fixes D::"'a \ 'b" + assumes "linear ( *) scale D" + obtains d where "D = (\x. scale x d)" +proof - + interpret linear "( *)" scale D by fact + show ?thesis + by (metis mult.commute mult.left_neutral scale that) +qed + +lemma scale_eq_0_iff [simp]: "scale a x = 0 \ a = 0 \ x = 0" + by (metis scale_left_commute right_inverse scale_one scale_scale scale_zero_left) + +lemma scale_left_imp_eq: + assumes nonzero: "a \ 0" + and scale: "scale a x = scale a y" + shows "x = y" +proof - + from scale have "scale a (x - y) = 0" + by (simp add: scale_right_diff_distrib) + with nonzero have "x - y = 0" by simp + then show "x = y" by (simp only: right_minus_eq) +qed + +lemma scale_right_imp_eq: + assumes nonzero: "x \ 0" + and scale: "scale a x = scale b x" + shows "a = b" +proof - + from scale have "scale (a - b) x = 0" + by (simp add: scale_left_diff_distrib) + with nonzero have "a - b = 0" by simp + then show "a = b" by (simp only: right_minus_eq) +qed + +lemma scale_cancel_left [simp]: "scale a x = scale a y \ x = y \ a = 0" + by (auto intro: scale_left_imp_eq) + +lemma scale_cancel_right [simp]: "scale a x = scale b x \ a = b \ x = 0" + by (auto intro: scale_right_imp_eq) + +lemma injective_scale: "c \ 0 \ inj (\x. scale c x)" + by (simp add: inj_on_def) + +lemma dependent_def: "dependent P \ (\a \ P. a \ span (P - {a}))" + unfolding dependent_explicit +proof safe + fix a assume aP: "a \ P" and "a \ span (P - {a})" + then obtain a S u + where aP: "a \ P" and fS: "finite S" and SP: "S \ P" "a \ S" and ua: "(\v\S. u v *s v) = a" + unfolding span_explicit by blast + let ?S = "insert a S" + let ?u = "\y. if y = a then - 1 else u y" + from fS SP have "(\v\?S. ?u v *s v) = 0" + by (simp add: if_distrib[of "\r. r *s a" for a] sum.If_cases field_simps Diff_eq[symmetric] ua) + moreover have "finite ?S" "?S \ P" "a \ ?S" "?u a \ 0" + using fS SP aP by auto + ultimately show "\t u. finite t \ t \ P \ (\v\t. u v *s v) = 0 \ (\v\t. u v \ 0)" by fast +next + fix S u v + assume fS: "finite S" and SP: "S \ P" and vS: "v \ S" + and uv: "u v \ 0" and u: "(\v\S. u v *s v) = 0" + let ?a = v + let ?S = "S - {v}" + let ?u = "\i. (- u i) / u v" + have th0: "?a \ P" "finite ?S" "?S \ P" + using fS SP vS by auto + have "(\v\?S. ?u v *s v) = (\v\S. (- (inverse (u ?a))) *s (u v *s v)) - ?u v *s v" + using fS vS uv by (simp add: sum_diff1 field_simps) + also have "\ = ?a" + unfolding scale_sum_right[symmetric] u using uv by simp + finally have "(\v\?S. ?u v *s v) = ?a" . + with th0 show "\a \ P. a \ span (P - {a})" + unfolding span_explicit by (auto intro!: bexI[where x="?a"] exI[where x="?S"] exI[where x="?u"]) +qed + +lemma dependent_single[simp]: "dependent {x} \ x = 0" + unfolding dependent_def by auto + +lemma in_span_insert: + assumes a: "a \ span (insert b S)" + and na: "a \ span S" + shows "b \ span (insert a S)" +proof - + from span_breakdown[of b "insert b S" a, OF insertI1 a] + obtain k where k: "a - k *s b \ span (S - {b})" by auto + have "k \ 0" + proof + assume "k = 0" + with k span_mono[of "S - {b}" S] have "a \ span S" by auto + with na show False by blast + qed + then have eq: "b = (1/k) *s a - (1/k) *s (a - k *s b)" + by (simp add: algebra_simps) + + from k have "(1/k) *s (a - k *s b) \ span (S - {b})" + by (rule span_scale) + also have "... \ span (insert a S)" + by (rule span_mono) auto + finally show ?thesis + using k by (subst eq) (blast intro: span_diff span_scale span_base) +qed + +lemma dependent_insertD: assumes a: "a \ span S" and S: "dependent (insert a S)" shows "dependent S" +proof - + have "a \ S" using a by (auto dest: span_base) + obtain b where b: "b = a \ b \ S" "b \ span (insert a S - {b})" + using S unfolding dependent_def by blast + have "b \ a" "b \ S" + using b \a \ S\ a by auto + with b have *: "b \ span (insert a (S - {b}))" + by (auto simp: insert_Diff_if) + show "dependent S" + proof cases + assume "b \ span (S - {b})" with \b \ S\ show ?thesis + by (auto simp add: dependent_def) + next + assume "b \ span (S - {b})" + with * have "a \ span (insert b (S - {b}))" by (rule in_span_insert) + with a show ?thesis + using \b \ S\ by (auto simp: insert_absorb) + qed +qed + +lemma independent_insertI: "a \ span S \ independent S \ independent (insert a S)" + by (auto dest: dependent_insertD) + +lemma independent_insert: + "independent (insert a S) \ (if a \ S then independent S else independent S \ a \ span S)" +proof - + have "a \ S \ a \ span S \ dependent (insert a S)" + by (auto simp: dependent_def) + then show ?thesis + by (auto intro: dependent_mono simp: independent_insertI) +qed + +lemma maximal_independent_subset_extend: + assumes "S \ V" "independent S" + shows "\B. S \ B \ B \ V \ independent B \ V \ span B" +proof - + let ?C = "{B. S \ B \ independent B \ B \ V}" + have "\M\?C. \X\?C. M \ X \ X = M" + proof (rule subset_Zorn) + fix C :: "'b set set" assume "subset.chain ?C C" + then have C: "\c. c \ C \ c \ V" "\c. c \ C \ S \ c" "\c. c \ C \ independent c" + "\c d. c \ C \ d \ C \ c \ d \ d \ c" + unfolding subset.chain_def by blast+ + + show "\U\?C. \X\C. X \ U" + proof cases + assume "C = {}" with assms show ?thesis + by (auto intro!: exI[of _ S]) + next + assume "C \ {}" + with C(2) have "S \ \C" + by auto + moreover have "independent (\C)" + by (intro independent_Union_directed C) + moreover have "\C \ V" + using C by auto + ultimately show ?thesis + by auto + qed + qed + then obtain B where B: "independent B" "B \ V" "S \ B" + and max: "\S. independent S \ S \ V \ B \ S \ S = B" + by auto + moreover + { assume "\ V \ span B" + then obtain v where "v \ V" "v \ span B" + by auto + with B have "independent (insert v B)" by (auto intro: dependent_insertD) + from max[OF this] \v \ V\ \B \ V\ + have "v \ B" + by auto + with \v \ span B\ have False + by (auto intro: span_base) } + ultimately show ?thesis + by (auto intro!: exI[of _ B]) +qed + +lemma maximal_independent_subset: "\B. B \ V \ independent B \ V \ span B" + by (metis maximal_independent_subset_extend[of "{}"] empty_subsetI independent_empty) + +text \Extends a basis from B to a basis of the entire space.\ +definition extend_basis :: "'b set \ 'b set" + where "extend_basis B = (SOME B'. B \ B' \ independent B' \ span B' = UNIV)" + +lemma + assumes B: "independent B" + shows extend_basis_superset: "B \ extend_basis B" + and independent_extend_basis: "independent (extend_basis B)" + and span_extend_basis[simp]: "span (extend_basis B) = UNIV" +proof - + define p where "p B' \ B \ B' \ independent B' \ span B' = UNIV" for B' + obtain B' where "p B'" + using maximal_independent_subset_extend[OF subset_UNIV B] by (auto simp: p_def) + then have "p (extend_basis B)" + unfolding extend_basis_def p_def[symmetric] by (rule someI) + then show "B \ extend_basis B" "independent (extend_basis B)" "span (extend_basis B) = UNIV" + by (auto simp: p_def) +qed + +lemma in_span_delete: + assumes a: "a \ span S" + and na: "a \ span (S - {b})" + shows "b \ span (insert a (S - {b}))" + apply (rule in_span_insert) + apply (rule set_rev_mp) + apply (rule a) + apply (rule span_mono) + apply blast + apply (rule na) + done + +lemma span_redundant: "x \ span S \ span (insert x S) = span S" + unfolding span_def by (rule hull_redundant) + +lemma span_trans: "x \ span S \ y \ span (insert x S) \ y \ span S" + by (simp only: span_redundant) + +lemma span_insert_0[simp]: "span (insert 0 S) = span S" + by (metis span_zero span_redundant) + +lemma span_delete_0 [simp]: "span(S - {0}) = span S" +proof + show "span (S - {0}) \ span S" + by (blast intro!: span_mono) +next + have "span S \ span(insert 0 (S - {0}))" + by (blast intro!: span_mono) + also have "... \ span(S - {0})" + using span_insert_0 by blast + finally show "span S \ span (S - {0})" . +qed + +lemma span_image_scale: + assumes "finite S" and nz: "\x. x \ S \ c x \ 0" + shows "span ((\x. c x *s x) ` S) = span S" +using assms +proof (induction S arbitrary: c) + case (empty c) show ?case by simp +next + case (insert x F c) + show ?case + proof (intro set_eqI iffI) + fix y + assume "y \ span ((\x. c x *s x) ` insert x F)" + then show "y \ span (insert x F)" + using insert by (force simp: span_breakdown_eq) + next + fix y + assume "y \ span (insert x F)" + then show "y \ span ((\x. c x *s x) ` insert x F)" + using insert + apply (clarsimp simp: span_breakdown_eq) + apply (rule_tac x="k / c x" in exI) + by simp + qed +qed + +lemma exchange_lemma: + assumes f: "finite t" + and i: "independent s" + and sp: "s \ span t" + shows "\t'. card t' = card t \ finite t' \ s \ t' \ t' \ s \ t \ s \ span t'" + using f i sp +proof (induct "card (t - s)" arbitrary: s t rule: less_induct) + case less + note ft = `finite t` and s = `independent s` and sp = `s \ span t` + let ?P = "\t'. card t' = card t \ finite t' \ s \ t' \ t' \ s \ t \ s \ span t'" + let ?ths = "\t'. ?P t'" + + { + assume st: "t \ s" + from spanning_subset_independent[OF st s sp] st ft span_mono[OF st] + have ?ths by (auto intro: span_base) + } + moreover + { + assume st:"\ t \ s" + from st obtain b where b: "b \ t" "b \ s" + by blast + from b have "t - {b} - s \ t - s" + by blast + then have cardlt: "card (t - {b} - s) < card (t - s)" + using ft by (auto intro: psubset_card_mono) + from b ft have ct0: "card t \ 0" + by auto + have ?ths + proof cases + assume stb: "s \ span (t - {b})" + from ft have ftb: "finite (t - {b})" + by auto + from less(1)[OF cardlt ftb s stb] + obtain u where u: "card u = card (t - {b})" "s \ u" "u \ s \ (t - {b})" "s \ span u" + and fu: "finite u" by blast + let ?w = "insert b u" + have th0: "s \ insert b u" + using u by blast + from u(3) b have "u \ s \ t" + by blast + then have th1: "insert b u \ s \ t" + using u b by blast + have bu: "b \ u" + using b u by blast + from u(1) ft b have "card u = (card t - 1)" + by auto + then have th2: "card (insert b u) = card t" + using card_insert_disjoint[OF fu bu] ct0 by auto + from u(4) have "s \ span u" . + also have "\ \ span (insert b u)" + by (rule span_mono) blast + finally have th3: "s \ span (insert b u)" . + from th0 th1 th2 th3 fu have th: "?P ?w" + by blast + from th show ?thesis by blast + next + assume stb: "\ s \ span (t - {b})" + from stb obtain a where a: "a \ s" "a \ span (t - {b})" + by blast + have ab: "a \ b" + using a b by blast + have at: "a \ t" + using a ab span_base[of a "t- {b}"] by auto + have mlt: "card ((insert a (t - {b})) - s) < card (t - s)" + using cardlt ft a b by auto + have ft': "finite (insert a (t - {b}))" + using ft by auto + { + fix x + assume xs: "x \ s" + have t: "t \ insert b (insert a (t - {b}))" + using b by auto + have bs: "b \ span (insert a (t - {b}))" + apply (rule in_span_delete) + using a sp unfolding subset_eq + apply auto + done + from xs sp have "x \ span t" + by blast + with span_mono[OF t] have x: "x \ span (insert b (insert a (t - {b})))" .. + from span_trans[OF bs x] have "x \ span (insert a (t - {b}))" . + } + then have sp': "s \ span (insert a (t - {b}))" + by blast + from less(1)[OF mlt ft' s sp'] obtain u where u: + "card u = card (insert a (t - {b}))" + "finite u" "s \ u" "u \ s \ insert a (t - {b})" + "s \ span u" by blast + from u a b ft at ct0 have "?P u" + by auto + then show ?thesis by blast + qed + } + ultimately show ?ths by blast +qed + +lemma independent_span_bound: + assumes f: "finite t" + and i: "independent s" + and sp: "s \ span t" + shows "finite s \ card s \ card t" + by (metis exchange_lemma[OF f i sp] finite_subset card_mono) + +lemma independent_explicit_finite_subsets: + "independent A \ (\S \ A. finite S \ (\u. (\v\S. u v *s v) = 0 \ (\v\S. u v = 0)))" + unfolding dependent_explicit [of A] by (simp add: disj_not2) + +lemma independent_if_scalars_zero: + assumes fin_A: "finite A" + and sum: "\f x. (\x\A. f x *s x) = 0 \ x \ A \ f x = 0" + shows "independent A" +proof (unfold independent_explicit_finite_subsets, clarify) + fix S v and u :: "'b \ 'a" + assume S: "S \ A" and v: "v \ S" + let ?g = "\x. if x \ S then u x else 0" + have "(\v\A. ?g v *s v) = (\v\S. u v *s v)" + using S fin_A by (auto intro!: sum.mono_neutral_cong_right) + also assume "(\v\S. u v *s v) = 0" + finally have "?g v = 0" using v S sum by force + thus "u v = 0" unfolding if_P[OF v] . +qed + +lemma bij_if_span_eq_span_bases: + assumes B: "independent B" and C: "independent C" + and eq: "span B = span C" + shows "\f. bij_betw f B C" +proof cases + assume "finite B \ finite C" + then have "finite B \ finite C \ card C = card B" + using independent_span_bound[of B C] independent_span_bound[of C B] assms + span_superset[of B] span_superset[of C] + by auto + then show ?thesis + by (auto intro!: finite_same_card_bij) +next + assume "\ (finite B \ finite C)" + then have "infinite B" "infinite C" by auto + { fix B C assume B: "independent B" and C: "independent C" and "infinite B" "infinite C" and eq: "span B = span C" + let ?R = "representation B" and ?R' = "representation C" let ?U = "\c. {v. ?R c v \ 0}" + have in_span_C [simp, intro]: \b \ B \ b \ span C\ for b unfolding eq[symmetric] by (rule span_base) + have in_span_B [simp, intro]: \c \ C \ c \ span B\ for c unfolding eq by (rule span_base) + have \B \ (\c\C. ?U c)\ + proof + fix b assume \b \ B\ + have \b \ span C\ + using \b \ B\ unfolding eq[symmetric] by (rule span_base) + have \(\v | ?R' b v \ 0. \w | ?R v w \ 0. (?R' b v * ?R v w) *s w) = + (\v | ?R' b v \ 0. ?R' b v *s (\w | ?R v w \ 0. ?R v w *s w))\ + by (simp add: scale_sum_right) + also have \\ = (\v | ?R' b v \ 0. ?R' b v *s v)\ + by (auto simp: sum_nonzero_representation_eq B eq span_base representation_ne_zero) + also have \\ = b\ + by (rule sum_nonzero_representation_eq[OF C \b \ span C\]) + finally have "?R b b = ?R (\v | ?R' b v \ 0. \w | ?R v w \ 0. (?R' b v * ?R v w) *s w) b" + by simp + also have \\ = (\v | ?R' b v \ 0. \w | ?R v w \ 0. ?R' b v * ?R v w * ?R w b)\ + using B \b \ B\ + apply (subst representation_sum[OF B]) + apply (fastforce intro: span_sum span_scale span_base representation_ne_zero) + apply (rule sum.cong[OF refl]) + apply (subst representation_sum[OF B]) + apply (simp add: span_sum span_scale span_base representation_ne_zero) + apply (simp add: representation_scale[OF B] span_base representation_ne_zero) + done + finally have "(\v | ?R' b v \ 0. \w | ?R v w \ 0. ?R' b v * ?R v w * ?R w b) \ 0" + using representation_basis[OF B \b \ B\] by auto + then obtain v w where bv: "?R' b v \ 0" and vw: "?R v w \ 0" and "?R' b v * ?R v w * ?R w b \ 0" + by (blast elim: sum.not_neutral_contains_not_neutral) + with representation_basis[OF B, of w] vw[THEN representation_ne_zero] + have \?R' b v \ 0\ \?R v b \ 0\ by (auto split: if_splits) + then show \b \ (\c\C. ?U c)\ + by (auto dest: representation_ne_zero) + qed + then have B_eq: \B = (\c\C. ?U c)\ + by (auto intro: span_base representation_ne_zero eq) + have "ordLeq3 (card_of B) (card_of C)" + proof (subst B_eq, rule card_of_UNION_ordLeq_infinite[OF \infinite C\]) + show "ordLeq3 (card_of C) (card_of C)" + by (intro ordLeq_refl card_of_Card_order) + show "\c\C. ordLeq3 (card_of {v. representation B c v \ 0}) (card_of C)" + by (intro ballI ordLeq3_finite_infinite \infinite C\ finite_representation) + qed } + from this[of B C] this[of C B] B C eq \infinite C\ \infinite B\ + show ?thesis by (auto simp add: ordIso_iff_ordLeq card_of_ordIso) +qed + +definition dim :: "'b set \ nat" + where "dim V = card (SOME b. independent b \ span b = span V)" + +lemma dim_eq_card: + assumes BV: "span B = span V" and B: "independent B" + shows "dim V = card B" +proof - + define p where "p b \ independent b \ span b = span V" for b + have "p (SOME B. p B)" + using assms by (intro someI[of p B]) (auto simp: p_def) + then have "\f. bij_betw f B (SOME B. p B)" + by (subst (asm) p_def, intro bij_if_span_eq_span_bases[OF B]) (simp_all add: BV) + then have "card B = card (SOME B. p B)" + by (auto intro: bij_betw_same_card) + then show ?thesis + by (simp add: dim_def p_def) +qed + +lemma basis_card_eq_dim: "B \ V \ V \ span B \ independent B \ card B = dim V" + using dim_eq_card[of B V] span_mono[of B V] span_minimal[OF _ subspace_span, of V B] by auto + +lemma basis_exists: "\B. B \ V \ independent B \ V \ span B \ card B = dim V" + by (meson basis_card_eq_dim empty_subsetI independent_empty maximal_independent_subset_extend) + +lemma dim_eq_card_independent: "independent B \ dim B = card B" + by (rule dim_eq_card[OF refl]) + +lemma dim_span[simp]: "dim (span S) = dim S" + by (simp add: dim_def span_span) + +lemma dim_span_eq_card_independent: "independent B \ dim (span B) = card B" + by (simp add: dim_span dim_eq_card) + +lemma dim_le_card: assumes "V \ span W" "finite W" shows "dim V \ card W" +proof - + obtain A where "independent A" "A \ V" "V \ span A" + using maximal_independent_subset[of V] by auto + with assms independent_span_bound[of W A] basis_card_eq_dim[of A V] + show ?thesis by auto +qed + +lemma span_eq_dim: "span S = span T \ dim S = dim T" + by (metis dim_span) + +corollary dim_le_card': + "finite s \ dim s \ card s" + by (metis basis_exists card_mono) + +lemma span_card_ge_dim: + "B \ V \ V \ span B \ finite B \ dim V \ card B" + by (simp add: dim_le_card) + +lemma dim_unique: + "B \ V \ V \ span B \ independent B \ card B = n \ dim V = n" + by (metis basis_card_eq_dim) + +lemma subspace_sums: "\subspace S; subspace T\ \ subspace {x + y|x y. x \ S \ y \ T}" + apply (simp add: subspace_def) + apply (intro conjI impI allI) + using add.right_neutral apply blast + apply clarify + apply (metis add.assoc add.left_commute) + using scale_right_distrib by blast + +end + +lemma linear_iff: "linear s1 s2 f \ + (vector_space s1 \ vector_space s2 \ (\x y. f (x + y) = f x + f y) \ (\c x. f (s1 c x) = s2 c (f x)))" + unfolding linear_def module_hom_iff vector_space_def module_def by auto + +context begin +qualified lemma linear_compose: "linear s1 s2 f \ linear s2 s3 g \ linear s1 s3 (g o f)" + unfolding module_hom_iff_linear[symmetric] + by (rule module_hom_compose) +end + +locale vector_space_pair = vs1: vector_space s1 + vs2: vector_space s2 + for s1 :: "'a::field \ 'b::ab_group_add \ 'b" (infixr "*a" 75) + and s2 :: "'a::field \ 'c::ab_group_add \ 'c" (infixr "*b" 75) +begin + +context fixes f assumes "linear s1 s2 f" begin +interpretation linear s1 s2 f by fact +lemmas\\from locale \module_hom\\ + linear_0 = zero + and linear_add = add + and linear_scale = scale + and linear_neg = neg + and linear_diff = diff + and linear_sum = sum + and linear_inj_on_iff_eq_0 = inj_on_iff_eq_0 + and linear_inj_iff_eq_0 = inj_iff_eq_0 + and linear_subspace_image = subspace_image + and linear_subspace_vimage = subspace_vimage + and linear_subspace_kernel = subspace_kernel + and linear_span_image = span_image + and linear_dependent_inj_imageD = dependent_inj_imageD + and linear_eq_0_on_span = eq_0_on_span + and linear_independent_injective_image = independent_injective_image + and linear_inj_on_span_independent_image = inj_on_span_independent_image + and linear_inj_on_span_iff_independent_image = inj_on_span_iff_independent_image + and linear_subspace_linear_preimage = subspace_linear_preimage + and linear_spans_image = spans_image + and linear_spanning_surjective_image = spanning_surjective_image +end + +sublocale module_pair + rewrites "module_hom = linear" + by unfold_locales (fact module_hom_eq_linear) + +lemmas\\from locale \module_pair\\ + linear_eq_on_span = module_hom_eq_on_span + and linear_compose_scale_right = module_hom_scale + and linear_compose_add = module_hom_add + and linear_zero = module_hom_zero + and linear_compose_sub = module_hom_sub + and linear_compose_neg = module_hom_neg + and linear_compose_scale = module_hom_compose_scale + +lemma linear_indep_image_lemma: + assumes lf: "linear s1 s2 f" + and fB: "finite B" + and ifB: "vs2.independent (f ` B)" + and fi: "inj_on f B" + and xsB: "x \ vs1.span B" + and fx: "f x = 0" + shows "x = 0" + using fB ifB fi xsB fx +proof (induct arbitrary: x rule: finite_induct[OF fB]) + case 1 + then show ?case by auto +next + case (2 a b x) + have fb: "finite b" using "2.prems" by simp + have th0: "f ` b \ f ` (insert a b)" + apply (rule image_mono) + apply blast + done + from vs2.independent_mono[ OF "2.prems"(2) th0] + have ifb: "vs2.independent (f ` b)" . + have fib: "inj_on f b" + apply (rule subset_inj_on [OF "2.prems"(3)]) + apply blast + done + from vs1.span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)] + obtain k where k: "x - k *a a \ vs1.span (b - {a})" + by blast + have "f (x - k *a a) \ vs2.span (f ` b)" + unfolding linear_span_image[OF lf] + apply (rule imageI) + using k vs1.span_mono[of "b - {a}" b] + apply blast + done + then have "f x - k *b f a \ vs2.span (f ` b)" + by (simp add: linear_diff linear_scale lf) + then have th: "-k *b f a \ vs2.span (f ` b)" + using "2.prems"(5) by simp + have xsb: "x \ vs1.span b" + proof (cases "k = 0") + case True + with k have "x \ vs1.span (b - {a})" by simp + then show ?thesis using vs1.span_mono[of "b - {a}" b] + by blast + next + case False + with vs2.span_scale[OF th, of "- 1/ k"] + have th1: "f a \ vs2.span (f ` b)" + by auto + from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric] + have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast + from "2.prems"(2) [unfolded vs2.dependent_def bex_simps(8), rule_format, of "f a"] + have "f a \ vs2.span (f ` b)" using tha + using "2.hyps"(2) + "2.prems"(3) by auto + with th1 have False by blast + then show ?thesis by blast + qed + from "2.hyps"(3)[OF fb ifb fib xsb "2.prems"(5)] show "x = 0" . +qed + +lemma linear_eq_on: + assumes l: "linear s1 s2 f" "linear s1 s2 g" + assumes x: "x \ vs1.span B" and eq: "\b. b \ B \ f b = g b" + shows "f x = g x" +proof - + interpret d: linear s1 s2 "\x. f x - g x" + using l by (intro linear_compose_sub) (auto simp: module_hom_iff_linear) + have "f x - g x = 0" + by (rule d.eq_0_on_span[OF _ x]) (auto simp: eq) + then show ?thesis by auto +qed + +definition construct :: "'b set \ ('b \ 'c) \ ('b \ 'c)" + where "construct B g v = (\b | vs1.representation (vs1.extend_basis B) v b \ 0. + vs1.representation (vs1.extend_basis B) v b *b (if b \ B then g b else 0))" + +lemma construct_cong: "(\b. b \ B \ f b = g b) \ construct B f = construct B g" + unfolding construct_def by (rule ext, auto intro!: sum.cong) + +lemma linear_construct: + assumes B[simp]: "vs1.independent B" + shows "linear s1 s2 (construct B f)" + unfolding module_hom_iff_linear linear_iff +proof safe + have eB[simp]: "vs1.independent (vs1.extend_basis B)" + using vs1.independent_extend_basis[OF B] . + let ?R = "vs1.representation (vs1.extend_basis B)" + fix c x y + have "construct B f (x + y) = + (\b\{b. ?R x b \ 0} \ {b. ?R y b \ 0}. ?R (x + y) b *b (if b \ B then f b else 0))" + by (auto intro!: sum.mono_neutral_cong_left simp: vs1.finite_representation vs1.representation_add construct_def) + also have "\ = construct B f x + construct B f y" + by (auto simp: construct_def vs1.representation_add vs2.scale_left_distrib sum.distrib + intro!: arg_cong2[where f="(+)"] sum.mono_neutral_cong_right vs1.finite_representation) + finally show "construct B f (x + y) = construct B f x + construct B f y" . + + show "construct B f (c *a x) = c *b construct B f x" + by (auto simp del: vs2.scale_scale intro!: sum.mono_neutral_cong_left vs1.finite_representation + simp add: construct_def vs2.scale_scale[symmetric] vs1.representation_scale vs2.scale_sum_right) +qed intro_locales + +lemma construct_basis: + assumes B[simp]: "vs1.independent B" and b: "b \ B" + shows "construct B f b = f b" +proof - + have *: "vs1.representation (vs1.extend_basis B) b = (\v. if v = b then 1 else 0)" + using vs1.extend_basis_superset[OF B] b + by (intro vs1.representation_basis vs1.independent_extend_basis) auto + then have "{v. vs1.representation (vs1.extend_basis B) b v \ 0} = {b}" + by auto + then show ?thesis + unfolding construct_def by (simp add: * b) +qed + +lemma construct_outside: + assumes B: "vs1.independent B" and v: "v \ vs1.span (vs1.extend_basis B - B)" + shows "construct B f v = 0" + unfolding construct_def +proof (clarsimp intro!: sum.neutral simp del: vs2.scale_eq_0_iff) + fix b assume "b \ B" + then have "vs1.representation (vs1.extend_basis B - B) v b = 0" + using vs1.representation_ne_zero[of "vs1.extend_basis B - B" v b] by auto + moreover have "vs1.representation (vs1.extend_basis B) v = vs1.representation (vs1.extend_basis B - B) v" + using vs1.representation_extend[OF vs1.independent_extend_basis[OF B] v] by auto + ultimately show "vs1.representation (vs1.extend_basis B) v b *b f b = 0" + by simp +qed + +lemma construct_add: + assumes B[simp]: "vs1.independent B" + shows "construct B (\x. f x + g x) v = construct B f v + construct B g v" +proof (rule linear_eq_on) + show "v \ vs1.span (vs1.extend_basis B)" by simp + show "b \ vs1.extend_basis B \ construct B (\x. f x + g x) b = construct B f b + construct B g b" for b + using construct_outside[OF B vs1.span_base, of b] by (cases "b \ B") (auto simp: construct_basis) +qed (intro linear_compose_add linear_construct B)+ + +lemma construct_scale: + assumes B[simp]: "vs1.independent B" + shows "construct B (\x. c *b f x) v = c *b construct B f v" +proof (rule linear_eq_on) + show "v \ vs1.span (vs1.extend_basis B)" by simp + show "b \ vs1.extend_basis B \ construct B (\x. c *b f x) b = c *b construct B f b" for b + using construct_outside[OF B vs1.span_base, of b] by (cases "b \ B") (auto simp: construct_basis) +qed (intro linear_construct module_hom_scale B)+ + +lemma construct_in_span: + assumes B[simp]: "vs1.independent B" + shows "construct B f v \ vs2.span (f ` B)" +proof - + interpret c: linear s1 s2 "construct B f" by (rule linear_construct) fact + let ?R = "vs1.representation B" + have "v \ vs1.span ((vs1.extend_basis B - B) \ B)" + by (auto simp: Un_absorb2 vs1.extend_basis_superset) + then obtain x y where "v = x + y" "x \ vs1.span (vs1.extend_basis B - B)" "y \ vs1.span B" + unfolding vs1.span_Un by auto + moreover have "construct B f (\b | ?R y b \ 0. ?R y b *a b) \ vs2.span (f ` B)" + by (auto simp add: c.sum c.scale construct_basis vs1.representation_ne_zero + intro!: vs2.span_sum vs2.span_scale intro: vs2.span_base ) + ultimately show "construct B f v \ vs2.span (f ` B)" + by (auto simp add: c.add construct_outside vs1.sum_nonzero_representation_eq) +qed + +lemma linear_compose_sum: + assumes lS: "\a \ S. linear s1 s2 (f a)" + shows "linear s1 s2 (\x. sum (\a. f a x) S)" +proof (cases "finite S") + case True + then show ?thesis + using lS by induct (simp_all add: linear_zero linear_compose_add) +next + case False + then show ?thesis + by (simp add: linear_zero) +qed + +lemma in_span_in_range_construct: + "x \ range (construct B f)" if i: "vs1.independent B" and x: "x \ vs2.span (f ` B)" +proof - + interpret linear "( *a)" "( *b)" "construct B f" + using i by (rule linear_construct) + obtain bb :: "('b \ 'c) \ ('b \ 'c) \ 'b set \ 'b" where + "\x0 x1 x2. (\v4. v4 \ x2 \ x1 v4 \ x0 v4) = (bb x0 x1 x2 \ x2 \ x1 (bb x0 x1 x2) \ x0 (bb x0 x1 x2))" + by moura + then have f2: "\B Ba f fa. (B \ Ba \ bb fa f Ba \ Ba \ f (bb fa f Ba) \ fa (bb fa f Ba)) \ f ` B = fa ` Ba" + by (meson image_cong) + have "vs1.span B \ vs1.span (vs1.extend_basis B)" + by (simp add: vs1.extend_basis_superset[OF i] vs1.span_mono) + then show "x \ range (construct B f)" + using f2 x by (metis (no_types) construct_basis[OF i, of _ f] + vs1.span_extend_basis[OF i] set_mp span_image spans_image) +qed + +lemma range_construct_eq_span: + "range (construct B f) = vs2.span (f ` B)" + if "vs1.independent B" + by (auto simp: that construct_in_span in_span_in_range_construct) + +lemma linear_independent_extend_subspace: + \\legacy: use @{term construct} instead\ + assumes "vs1.independent B" + shows "\g. linear s1 s2 g \ (\x\B. g x = f x) \ range g = vs2.span (f`B)" + by (rule exI[where x="construct B f"]) + (auto simp: linear_construct assms construct_basis range_construct_eq_span) + +lemma linear_independent_extend: + "vs1.independent B \ \g. linear s1 s2 g \ (\x\B. g x = f x)" + using linear_independent_extend_subspace[of B f] by auto + +lemma linear_exists_left_inverse_on: + assumes lf: "linear s1 s2 f" + assumes V: "vs1.subspace V" and f: "inj_on f V" + shows "\g\UNIV \ V. linear s2 s1 g \ (\v\V. g (f v) = v)" +proof - + interpret linear s1 s2 f by fact + obtain B where V_eq: "V = vs1.span B" and B: "vs1.independent B" + using vs1.maximal_independent_subset[of V] vs1.span_minimal[OF _ \vs1.subspace V\] by auto + have f: "inj_on f (vs1.span B)" + using f unfolding V_eq . + show ?thesis + proof (intro bexI ballI conjI) + interpret p: vector_space_pair s2 s1 by unfold_locales + have fB: "vs2.independent (f ` B)" + using independent_injective_image[OF B f] . + let ?g = "p.construct (f ` B) (the_inv_into B f)" + show "linear ( *b) ( *a) ?g" + by (rule p.linear_construct[OF fB]) + have "?g b \ vs1.span (the_inv_into B f ` f ` B)" for b + by (intro p.construct_in_span fB) + moreover have "the_inv_into B f ` f ` B = B" + by (auto simp: image_comp comp_def the_inv_into_f_f inj_on_subset[OF f vs1.span_superset] + cong: image_cong) + ultimately show "?g \ UNIV \ V" + by (auto simp: V_eq) + have "(?g \ f) v = id v" if "v \ vs1.span B" for v + proof (rule vector_space_pair.linear_eq_on[where x=v]) + show "vector_space_pair ( *a) ( *a)" by unfold_locales + show "linear ( *a) ( *a) (?g \ f)" + apply (rule Vector_Spaces.linear_compose[of _ "( *b)"]) + subgoal by unfold_locales + apply fact + done + show "linear ( *a) ( *a) id" by (rule vs1.linear_id) + show "v \ vs1.span B" by fact + show "b \ B \ (p.construct (f ` B) (the_inv_into B f) \ f) b = id b" for b + by (simp add: p.construct_basis fB the_inv_into_f_f inj_on_subset[OF f vs1.span_superset]) + qed + then show "v \ V \ ?g (f v) = v" for v by (auto simp: comp_def id_def V_eq) + qed +qed + +lemma linear_exists_right_inverse_on: + assumes lf: "linear s1 s2 f" + assumes "vs1.subspace V" + shows "\g\UNIV \ V. linear s2 s1 g \ (\v\f ` V. f (g v) = v)" +proof - + obtain B where V_eq: "V = vs1.span B" and B: "vs1.independent B" + using vs1.maximal_independent_subset[of V] vs1.span_minimal[OF _ \vs1.subspace V\] by auto + obtain C where C: "vs2.independent C" and fB_C: "f ` B \ vs2.span C" "C \ f ` B" + using vs2.maximal_independent_subset[of "f ` B"] by auto + then have "\v\C. \b\B. v = f b" by auto + then obtain g where g: "\v. v \ C \ g v \ B" "\v. v \ C \ f (g v) = v" by metis + show ?thesis + proof (intro bexI ballI conjI) + interpret p: vector_space_pair s2 s1 by unfold_locales + let ?g = "p.construct C g" + show "linear ( *b) ( *a) ?g" + by (rule p.linear_construct[OF C]) + have "?g v \ vs1.span (g ` C)" for v + by (rule p.construct_in_span[OF C]) + also have "\ \ V" unfolding V_eq using g by (intro vs1.span_mono) auto + finally show "?g \ UNIV \ V" by auto + have "(f \ ?g) v = id v" if v: "v \ f ` V" for v + proof (rule vector_space_pair.linear_eq_on[where x=v]) + show "vector_space_pair ( *b) ( *b)" by unfold_locales + show "linear ( *b) ( *b) (f \ ?g)" + apply (rule Vector_Spaces.linear_compose[of _ "( *a)"]) + apply fact + subgoal by fact + done + show "linear ( *b) ( *b) id" by (rule vs2.linear_id) + have "vs2.span (f ` B) = vs2.span C" + using fB_C vs2.span_mono[of C "f ` B"] vs2.span_minimal[of "f`B" "vs2.span C"] by (auto simp: vs2.subspace_span) + then show "v \ vs2.span C" + using v linear_span_image[OF lf, of B] by (simp add: V_eq) + show "(f \ p.construct C g) b = id b" if b: "b \ C" for b + by (auto simp: p.construct_basis g C b) + qed + then show "v \ f ` V \ f (?g v) = v" for v by (auto simp: comp_def id_def) + qed +qed + +lemma linear_inj_on_left_inverse: + assumes lf: "linear s1 s2 f" + assumes fi: "inj_on f (vs1.span S)" + shows "\g. range g \ vs1.span S \ linear s2 s1 g \ (\x\vs1.span S. g (f x) = x)" + using linear_exists_left_inverse_on[OF lf vs1.subspace_span fi] + by (auto simp: linear_iff_module_hom) + +lemma linear_injective_left_inverse: "linear s1 s2 f \ inj f \ \g. linear s2 s1 g \ g \ f = id" + using linear_inj_on_left_inverse[of f UNIV] by (auto simp: fun_eq_iff vs1.span_UNIV) + +lemma linear_surj_right_inverse: + assumes lf: "linear s1 s2 f" + assumes sf: "vs2.span T \ f`vs1.span S" + shows "\g. range g \ vs1.span S \ linear s2 s1 g \ (\x\vs2.span T. f (g x) = x)" + using linear_exists_right_inverse_on[OF lf vs1.subspace_span, of S] sf + by (auto simp: linear_iff_module_hom) + +lemma linear_surjective_right_inverse: "linear s1 s2 f \ surj f \ \g. linear s2 s1 g \ f \ g = id" + using linear_surj_right_inverse[of f UNIV UNIV] + by (auto simp: vs1.span_UNIV vs2.span_UNIV fun_eq_iff) + +end + +lemma surjective_iff_injective_gen: + assumes fS: "finite S" + and fT: "finite T" + and c: "card S = card T" + and ST: "f ` S \ T" + shows "(\y \ T. \x \ S. f x = y) \ inj_on f S" + (is "?lhs \ ?rhs") +proof + assume h: "?lhs" + { + fix x y + assume x: "x \ S" + assume y: "y \ S" + assume f: "f x = f y" + from x fS have S0: "card S \ 0" + by auto + have "x = y" + proof (rule ccontr) + assume xy: "\ ?thesis" + have th: "card S \ card (f ` (S - {y}))" + unfolding c + apply (rule card_mono) + apply (rule finite_imageI) + using fS apply simp + using h xy x y f unfolding subset_eq image_iff + apply auto + apply (case_tac "xa = f x") + apply (rule bexI[where x=x]) + apply auto + done + also have " \ \ card (S - {y})" + apply (rule card_image_le) + using fS by simp + also have "\ \ card S - 1" using y fS by simp + finally show False using S0 by arith + qed + } + then show ?rhs + unfolding inj_on_def by blast +next + assume h: ?rhs + have "f ` S = T" + apply (rule card_subset_eq[OF fT ST]) + unfolding card_image[OF h] + apply (rule c) + done + then show ?lhs by blast +qed + +locale finite_dimensional_vector_space = vector_space + + fixes Basis :: "'b set" + assumes finite_Basis: "finite Basis" + and independent_Basis: "independent Basis" + and span_Basis: "span Basis = UNIV" +begin + +definition "dimension = card Basis" + +lemma finiteI_independent: "independent B \ finite B" + using independent_span_bound[OF finite_Basis, of B] by (auto simp: span_Basis) + +lemma dim_empty [simp]: "dim {} = 0" + by (rule dim_unique[OF order_refl]) (auto simp: dependent_def) + +lemma dim_insert: + "dim (insert x S) = (if x \ span S then dim S else dim S + 1)" +proof - + show ?thesis + proof (cases "x \ span S") + case True then show ?thesis + by (metis dim_span span_redundant) + next + case False + obtain B where B: "B \ span S" "independent B" "span S \ span B" "card B = dim (span S)" + using basis_exists [of "span S"] by blast + have 1: "insert x B \ span (insert x S)" + by (meson B(1) insertI1 insert_subset order_trans span_base span_mono subset_insertI) + have 2: "span (insert x S) \ span (insert x B)" + by (metis \B \ span S\ \span S \ span B\ span_breakdown_eq span_subspace subsetI subspace_span) + have 3: "independent (insert x B)" + by (metis B independent_insert span_subspace subspace_span False) + have "dim (span (insert x S)) = Suc (dim S)" + apply (rule dim_unique [OF 1 2 3]) + by (metis B False card_insert_disjoint dim_span finiteI_independent span_base span_eq span_span) + then show ?thesis + by (metis False Suc_eq_plus1 dim_span) + qed +qed + +lemma dim_singleton [simp]: + "dim{x} = (if x = 0 then 0 else 1)" + by (simp add: dim_insert) + +proposition choose_subspace_of_subspace: + assumes "n \ dim S" + obtains T where "subspace T" "T \ span S" "dim T = n" +proof - + have "\T. subspace T \ T \ span S \ dim T = n" + using assms + proof (induction n) + case 0 then show ?case by (auto intro!: exI[where x="{0}"] span_zero) + next + case (Suc n) + then obtain T where "subspace T" "T \ span S" "dim T = n" + by force + then show ?case + proof (cases "span S \ span T") + case True + have "dim S = dim T" + apply (rule span_eq_dim [OF subset_antisym [OF True]]) + by (simp add: \T \ span S\ span_minimal subspace_span) + then show ?thesis + using Suc.prems \dim T = n\ by linarith + next + case False + then obtain y where y: "y \ S" "y \ T" + by (meson span_mono subsetI) + then have "span (insert y T) \ span S" + by (metis (no_types) \T \ span S\ subsetD insert_subset span_superset span_mono span_span) + with \dim T = n\ \subspace T\ y show ?thesis + apply (rule_tac x="span(insert y T)" in exI) + apply (auto simp: dim_insert dim_span subspace_span) + using span_eq_iff by blast + qed + qed + with that show ?thesis by blast +qed + +lemma basis_subspace_exists: + "subspace S + \ \b. finite b \ b \ S \ + independent b \ span b = S \ card b = dim S" + by (meson basis_exists finiteI_independent span_subspace) + +lemma dim_mono: assumes "V \ span W" shows "dim V \ dim W" +proof - + obtain B where "independent B" "B \ W" "W \ span B" + using maximal_independent_subset[of W] by auto + with dim_le_card[of V B] assms independent_span_bound[of Basis B] basis_card_eq_dim[of B W] + span_mono[of B W] span_minimal[OF _ subspace_span, of W B] + show ?thesis + by (auto simp: finite_Basis span_Basis) +qed + +lemma dim_subset: "S \ T \ dim S \ dim T" + using dim_mono[of S T] by (auto intro: span_base) + +lemma dim_eq_0 [simp]: + "dim S = 0 \ S \ {0}" + using basis_exists finiteI_independent + apply safe + subgoal by fastforce + by (metis dim_singleton dim_subset le_0_eq) + +lemma dim_UNIV[simp]: "dim UNIV = card Basis" + using dim_eq_card[of Basis UNIV] by (simp add: independent_Basis span_Basis span_UNIV) + +lemma independent_card_le_dim: assumes "B \ V" and "independent B" shows "card B \ dim V" + by (subst dim_eq_card[symmetric, OF refl \independent B\]) (rule dim_subset[OF \B \ V\]) + +lemma dim_subset_UNIV: "dim S \ dimension" + by (metis dim_subset subset_UNIV dim_UNIV dimension_def) + +lemma card_ge_dim_independent: + assumes BV: "B \ V" + and iB: "independent B" + and dVB: "dim V \ card B" + shows "V \ span B" +proof + fix a + assume aV: "a \ V" + { + assume aB: "a \ span B" + then have iaB: "independent (insert a B)" + using iB aV BV by (simp add: independent_insert) + from aV BV have th0: "insert a B \ V" + by blast + from aB have "a \B" + by (auto simp add: span_base) + with independent_card_le_dim[OF th0 iaB] dVB finiteI_independent[OF iB] + have False by auto + } + then show "a \ span B" by blast +qed + +lemma card_le_dim_spanning: + assumes BV: "B \ V" + and VB: "V \ span B" + and fB: "finite B" + and dVB: "dim V \ card B" + shows "independent B" +proof - + { + fix a + assume a: "a \ B" "a \ span (B - {a})" + from a fB have c0: "card B \ 0" + by auto + from a fB have cb: "card (B - {a}) = card B - 1" + by auto + { + fix x + assume x: "x \ V" + from a have eq: "insert a (B - {a}) = B" + by blast + from x VB have x': "x \ span B" + by blast + from span_trans[OF a(2), unfolded eq, OF x'] + have "x \ span (B - {a})" . + } + then have th1: "V \ span (B - {a})" + by blast + have th2: "finite (B - {a})" + using fB by auto + from dim_le_card[OF th1 th2] + have c: "dim V \ card (B - {a})" . + from c c0 dVB cb have False by simp + } + then show ?thesis + unfolding dependent_def by blast +qed + +lemma card_eq_dim: "B \ V \ card B = dim V \ finite B \ independent B \ V \ span B" + by (metis order_eq_iff card_le_dim_spanning card_ge_dim_independent) + +lemma subspace_dim_equal: + assumes "subspace S" + and "subspace T" + and "S \ T" + and "dim S \ dim T" + shows "S = T" +proof - + obtain B where B: "B \ S" "independent B \ S \ span B" "card B = dim S" + using basis_exists[of S] by auto + then have "span B \ S" + using span_mono[of B S] span_eq_iff[of S] assms by metis + then have "span B = S" + using B by auto + have "dim S = dim T" + using assms dim_subset[of S T] by auto + then have "T \ span B" + using card_eq_dim[of B T] B finiteI_independent assms by auto + then show ?thesis + using assms \span B = S\ by auto +qed + +corollary dim_eq_span: + shows "\S \ T; dim T \ dim S\ \ span S = span T" + by (simp add: dim_span span_mono subspace_dim_equal subspace_span) + +lemma dim_psubset: + "span S \ span T \ dim S < dim T" + by (metis (no_types, hide_lams) dim_span less_le not_le subspace_dim_equal subspace_span) + +lemma dim_eq_full: + shows "dim S = dimension \ span S = UNIV" + by (metis dim_eq_span dim_subset_UNIV span_Basis span_span subset_UNIV + dim_UNIV dim_span dimension_def) + +lemma indep_card_eq_dim_span: + assumes "independent B" + shows "finite B \ card B = dim (span B)" + using dim_span_eq_card_independent[OF assms] finiteI_independent[OF assms] by auto + +text \More general size bound lemmas.\ + +lemma independent_bound_general: + "independent S \ finite S \ card S \ dim S" + by (simp add: dim_eq_card_independent finiteI_independent) + +lemma independent_explicit: + shows "independent B \ finite B \ (\c. (\v\B. c v *s v) = 0 \ (\v \ B. c v = 0))" + apply (cases "finite B") + apply (force simp: dependent_finite) + using independent_bound_general + apply auto + done + +proposition dim_sums_Int: + assumes "subspace S" "subspace T" + shows "dim {x + y |x y. x \ S \ y \ T} + dim(S \ T) = dim S + dim T" +proof - + obtain B where B: "B \ S \ T" "S \ T \ span B" + and indB: "independent B" + and cardB: "card B = dim (S \ T)" + using basis_exists by blast + then obtain C D where "B \ C" "C \ S" "independent C" "S \ span C" + and "B \ D" "D \ T" "independent D" "T \ span D" + using maximal_independent_subset_extend + by (metis Int_subset_iff \B \ S \ T\ indB) + then have "finite B" "finite C" "finite D" + by (simp_all add: finiteI_independent indB independent_bound_general) + have Beq: "B = C \ D" + apply (rule sym) + apply (rule spanning_subset_independent) + using \B \ C\ \B \ D\ apply blast + apply (meson \independent C\ independent_mono inf.cobounded1) + using B \C \ S\ \D \ T\ apply auto + done + then have Deq: "D = B \ (D - C)" + by blast + have CUD: "C \ D \ {x + y |x y. x \ S \ y \ T}" + apply safe + apply (metis add.right_neutral subsetCE \C \ S\ \subspace T\ set_eq_subset span_zero span_minimal) + apply (metis add.left_neutral subsetCE \D \ T\ \subspace S\ set_eq_subset span_zero span_minimal) + done + have "a v = 0" if 0: "(\v\C. a v *s v) + (\v\D - C. a v *s v) = 0" + and v: "v \ C \ (D-C)" for a v + proof - + have eq: "(\v\D - C. a v *s v) = - (\v\C. a v *s v)" + using that add_eq_0_iff by blast + have "(\v\D - C. a v *s v) \ S" + apply (subst eq) + apply (rule subspace_neg [OF \subspace S\]) + apply (rule subspace_sum [OF \subspace S\]) + by (meson subsetCE subspace_scale \C \ S\ \subspace S\) + moreover have "(\v\D - C. a v *s v) \ T" + apply (rule subspace_sum [OF \subspace T\]) + by (meson DiffD1 \D \ T\ \subspace T\ subset_eq subspace_def) + ultimately have "(\v \ D-C. a v *s v) \ span B" + using B by blast + then obtain e where e: "(\v\B. e v *s v) = (\v \ D-C. a v *s v)" + using span_finite [OF \finite B\] by force + have "\c v. \(\v\C. c v *s v) = 0; v \ C\ \ c v = 0" + using \finite C\ \independent C\ independentD by blast + define cc where "cc x = (if x \ B then a x + e x else a x)" for x + have [simp]: "C \ B = B" "D \ B = B" "C \ - B = C-D" "B \ (D - C) = {}" + using \B \ C\ \B \ D\ Beq by blast+ + have f2: "(\v\C \ D. e v *s v) = (\v\D - C. a v *s v)" + using Beq e by presburger + have f3: "(\v\C \ D. a v *s v) = (\v\C - D. a v *s v) + (\v\D - C. a v *s v) + (\v\C \ D. a v *s v)" + using \finite C\ \finite D\ sum.union_diff2 by blast + have f4: "(\v\C \ (D - C). a v *s v) = (\v\C. a v *s v) + (\v\D - C. a v *s v)" + by (meson Diff_disjoint \finite C\ \finite D\ finite_Diff sum.union_disjoint) + have "(\v\C. cc v *s v) = 0" + using 0 f2 f3 f4 + apply (simp add: cc_def Beq \finite C\ sum.If_cases algebra_simps sum.distrib + if_distrib if_distribR) + apply (simp add: add.commute add.left_commute diff_eq) + done + then have "\v. v \ C \ cc v = 0" + using independent_explicit \independent C\ \finite C\ by blast + then have C0: "\v. v \ C - B \ a v = 0" + by (simp add: cc_def Beq) meson + then have [simp]: "(\x\C - B. a x *s x) = 0" + by simp + have "(\x\C. a x *s x) = (\x\B. a x *s x)" + proof - + have "C - D = C - B" + using Beq by blast + then show ?thesis + using Beq \(\x\C - B. a x *s x) = 0\ f3 f4 by auto + qed + with 0 have Dcc0: "(\v\D. a v *s v) = 0" + apply (subst Deq) + by (simp add: \finite B\ \finite D\ sum_Un) + then have D0: "\v. v \ D \ a v = 0" + using independent_explicit \independent D\ \finite D\ by blast + show ?thesis + using v C0 D0 Beq by blast + qed + then have "independent (C \ (D - C))" + unfolding independent_explicit + using independent_explicit + by (simp add: independent_explicit \finite C\ \finite D\ sum_Un del: Un_Diff_cancel) + then have indCUD: "independent (C \ D)" by simp + have "dim (S \ T) = card B" + by (rule dim_unique [OF B indB refl]) + moreover have "dim S = card C" + by (metis \C \ S\ \independent C\ \S \ span C\ basis_card_eq_dim) + moreover have "dim T = card D" + by (metis \D \ T\ \independent D\ \T \ span D\ basis_card_eq_dim) + moreover have "dim {x + y |x y. x \ S \ y \ T} = card(C \ D)" + apply (rule dim_unique [OF CUD _ indCUD refl], clarify) + apply (meson \S \ span C\ \T \ span D\ span_add span_superset span_minimal subsetCE subspace_span sup.bounded_iff) + done + ultimately show ?thesis + using \B = C \ D\ [symmetric] + by (simp add: \independent C\ \independent D\ card_Un_Int finiteI_independent) +qed + +lemma dependent_biggerset_general: + "(finite S \ card S > dim S) \ dependent S" + using independent_bound_general[of S] by (metis linorder_not_le) + +lemma subset_le_dim: + "S \ span T \ dim S \ dim T" + by (metis dim_span dim_subset) + +lemma linear_inj_imp_surj: + assumes lf: "linear scale scale f" + and fi: "inj f" + shows "surj f" +proof - + interpret lf: linear scale scale f by fact + from basis_exists[of UNIV] obtain B + where B: "B \ UNIV" "independent B" "UNIV \ span B" "card B = dim UNIV" + by blast + from B(4) have d: "dim UNIV = card B" + by simp + have th: "UNIV \ span (f ` B)" + apply (rule card_ge_dim_independent) + apply blast + using B(2) inj_on_subset[OF fi] + apply (rule lf.independent_inj_on_image) + apply blast + apply (rule order_eq_refl) + apply (rule sym) + unfolding d + apply (rule card_image) + apply (rule subset_inj_on[OF fi]) + apply blast + done + from th show ?thesis + unfolding lf.span_image surj_def + using B(3) by blast +qed + +end + +locale finite_dimensional_vector_space_pair = + vs1: finite_dimensional_vector_space s1 B1 + vs2: finite_dimensional_vector_space s2 B2 + for s1 :: "'a::field \ 'b::ab_group_add \ 'b" (infixr "*a" 75) + and B1 :: "'b set" + and s2 :: "'a::field \ 'c::ab_group_add \ 'c" (infixr "*b" 75) + and B2 :: "'c set" +begin + +sublocale vector_space_pair s1 s2 by unfold_locales + +lemma linear_surjective_imp_injective: + assumes lf: "linear s1 s2 f" and sf: "surj f" and eq: "vs2.dim UNIV = vs1.dim UNIV" + shows "inj f" +proof - + interpret linear s1 s2 f by fact + have *: "card (f ` B1) \ vs2.dim UNIV" + using vs1.finite_Basis vs1.dim_eq_card[of B1 UNIV] sf + by (auto simp: vs1.span_Basis vs1.span_UNIV vs1.independent_Basis eq + simp del: vs2.dim_UNIV + intro!: card_image_le) + have indep_fB: "vs2.independent (f ` B1)" + using vs1.finite_Basis vs1.dim_eq_card[of B1 UNIV] sf * + by (intro vs2.card_le_dim_spanning[of "f ` B1" UNIV]) (auto simp: span_image vs1.span_Basis ) + have "vs2.dim UNIV \ card (f ` B1)" + unfolding eq sf[symmetric] vs2.dim_span_eq_card_independent[symmetric, OF indep_fB] + vs2.dim_span + by (intro vs2.dim_mono) (auto simp: span_image vs1.span_Basis) + with * have "card (f ` B1) = vs2.dim UNIV" by auto + also have "... = card B1" + unfolding eq vs1.dim_UNIV .. + finally have "inj_on f B1" + by (subst inj_on_iff_eq_card[OF vs1.finite_Basis]) + then show "inj f" + using inj_on_span_iff_independent_image[OF indep_fB] vs1.span_Basis by auto +qed + +lemma linear_injective_imp_surjective: + assumes lf: "linear s1 s2 f" and sf: "inj f" and eq: "vs2.dim UNIV = vs1.dim UNIV" + shows "surj f" +proof - + interpret linear s1 s2 f by fact + have *: False if b: "b \ vs2.span (f ` B1)" for b + proof - + have *: "vs2.independent (f ` B1)" + using vs1.independent_Basis by (intro independent_injective_image inj_on_subset[OF sf]) auto + have **: "vs2.independent (insert b (f ` B1))" + using b * by (rule vs2.independent_insertI) + + have "b \ f ` B1" using vs2.span_base[of b "f ` B1"] b by auto + then have "Suc (card B1) = card (insert b (f ` B1))" + using sf[THEN inj_on_subset, of B1] by (subst card_insert) (auto intro: vs1.finite_Basis simp: card_image) + also have "\ = vs2.dim (insert b (f ` B1))" + using vs2.dim_eq_card_independent[OF **] by simp + also have "vs2.dim (insert b (f ` B1)) \ vs2.dim B2" + by (rule vs2.dim_mono) (auto simp: vs2.span_Basis) + also have "\ = card B1" + using vs1.dim_span[of B1] vs2.dim_span[of B2] unfolding vs1.span_Basis vs2.span_Basis eq + vs1.dim_eq_card_independent[OF vs1.independent_Basis] by simp + finally show False by simp + qed + have "f ` UNIV = f ` vs1.span B1" unfolding vs1.span_Basis .. + also have "\ = vs2.span (f ` B1)" unfolding span_image .. + also have "\ = UNIV" using * by blast + finally show ?thesis . +qed + +lemma linear_injective_isomorphism: + assumes lf: "linear s1 s2 f" + and fi: "inj f" + and dims: "vs2.dim UNIV = vs1.dim UNIV" + shows "\f'. linear s2 s1 f' \ (\x. f' (f x) = x) \ (\x. f (f' x) = x)" +proof - + show ?thesis + unfolding isomorphism_expand[symmetric] + using linear_exists_right_inverse_on[OF lf vs1.subspace_UNIV] + linear_exists_left_inverse_on[OF lf vs1.subspace_UNIV fi] + apply (auto simp: module_hom_iff_linear) + subgoal for f' f'' + apply (rule exI[where x=f'']) + using linear_injective_imp_surjective[OF lf fi dims] + apply auto + by (metis comp_apply eq_id_iff surj_def) + done +qed + +lemma linear_surjective_isomorphism: + assumes lf: "linear s1 s2 f" + and sf: "surj f" + and dims: "vs2.dim UNIV = vs1.dim UNIV" + shows "\f'. linear s2 s1 f' \ (\x. f' (f x) = x) \ (\x. f (f' x) = x)" +proof - + show ?thesis + unfolding isomorphism_expand[symmetric] + using linear_exists_right_inverse_on[OF lf vs1.subspace_UNIV] + linear_exists_left_inverse_on[OF lf vs1.subspace_UNIV] + using linear_surjective_imp_injective[OF lf sf dims] sf + apply (auto simp: module_hom_iff_linear) + subgoal for f' f'' + apply (rule exI[where x=f'']) + apply auto + by (metis isomorphism_expand) + done +qed + +lemma dim_image_eq: + assumes lf: "linear s1 s2 f" + and fi: "inj_on f (vs1.span S)" + shows "vs2.dim (f ` S) = vs1.dim S" +proof - + interpret lf: linear by fact + obtain B where B: "B \ S" "vs1.independent B" "S \ vs1.span B" "card B = vs1.dim S" + using vs1.basis_exists[of S] by auto + then have "vs1.span S = vs1.span B" + using vs1.span_mono[of B S] vs1.span_mono[of S "vs1.span B"] vs1.span_span[of B] by auto + moreover have "card (f ` B) = card B" + using assms card_image[of f B] subset_inj_on[of f "vs1.span S" B] B vs1.span_superset by auto + moreover have "(f ` B) \ (f ` S)" + using B by auto + ultimately show ?thesis + by (metis B(2) B(4) fi lf.dependent_inj_imageD lf.span_image vs2.dim_eq_card_independent vs2.dim_span) +qed + +lemma basis_to_basis_subspace_isomorphism: + assumes s: "vs1.subspace S" + and t: "vs2.subspace T" + and d: "vs1.dim S = vs2.dim T" + and B: "B \ S" "vs1.independent B" "S \ vs1.span B" "card B = vs1.dim S" + and C: "C \ T" "vs2.independent C" "T \ vs2.span C" "card C = vs2.dim T" + shows "\f. linear s1 s2 f \ f ` B = C \ f ` S = T \ inj_on f S" +proof - + from B have fB: "finite B" + by (simp add: vs1.finiteI_independent) + from C have fC: "finite C" + by (simp add: vs2.finiteI_independent) + from B(4) C(4) card_le_inj[of B C] d obtain f where + f: "f ` B \ C" "inj_on f B" using \finite B\ \finite C\ by auto + from linear_independent_extend[OF B(2)] obtain g where + g: "linear s1 s2 g" "\x \ B. g x = f x" by blast + interpret g: linear s1 s2 g by fact + from inj_on_iff_eq_card[OF fB, of f] f(2) + have "card (f ` B) = card B" by simp + with B(4) C(4) have ceq: "card (f ` B) = card C" using d + by simp + have "g ` B = f ` B" using g(2) + by (auto simp add: image_iff) + also have "\ = C" using card_subset_eq[OF fC f(1) ceq] . + finally have gBC: "g ` B = C" . + have gi: "inj_on g B" using f(2) g(2) + by (auto simp add: inj_on_def) + note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] + { + fix x y + assume x: "x \ S" and y: "y \ S" and gxy: "g x = g y" + from B(3) x y have x': "x \ vs1.span B" and y': "y \ vs1.span B" + by blast+ + from gxy have th0: "g (x - y) = 0" + by (simp add: g.diff) + have th1: "x - y \ vs1.span B" using x' y' + by (metis vs1.span_diff) + have "x = y" using g0[OF th1 th0] by simp + } + then have giS: "inj_on g S" unfolding inj_on_def by blast + from vs1.span_subspace[OF B(1,3) s] + have "g ` S = vs2.span (g ` B)" + by (simp add: g.span_image) + also have "\ = vs2.span C" + unfolding gBC .. + also have "\ = T" + using vs2.span_subspace[OF C(1,3) t] . + finally have gS: "g ` S = T" . + from g(1) gS giS gBC show ?thesis + by blast +qed + +lemma dim_image_le: + assumes lf: "linear s1 s2 f" + shows "vs2.dim (f ` S) \ vs1.dim (S)" +proof - + from vs1.basis_exists[of S] obtain B where + B: "B \ S" "vs1.independent B" "S \ vs1.span B" "card B = vs1.dim S" by blast + from B have fB: "finite B" "card B = vs1.dim S" + using vs1.independent_bound_general by blast+ + have "vs2.dim (f ` S) \ card (f ` B)" + apply (rule vs2.span_card_ge_dim) + using lf B fB + apply (auto simp add: module_hom.span_image module_hom.spans_image subset_image_iff + linear_iff_module_hom) + done + also have "\ \ vs1.dim S" + using card_image_le[OF fB(1)] fB by simp + finally show ?thesis . +qed + +end + +context finite_dimensional_vector_space begin + +lemma linear_surj_imp_inj: + assumes lf: "linear scale scale f" + and sf: "surj f" + shows "inj f" +proof - + interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales + let ?U = "UNIV :: 'b set" + from basis_exists[of ?U] obtain B + where B: "B \ ?U" "independent B" "?U \ span B" and d: "card B = dim ?U" + by blast + { + fix x + assume x: "x \ span B" + assume fx: "f x = 0" + from B(2) have fB: "finite B" + using finiteI_independent by auto + have fBi: "independent (f ` B)" + apply (rule card_le_dim_spanning[of "f ` B" ?U]) + apply blast + using sf B(3) + unfolding linear_span_image[OF lf] surj_def subset_eq image_iff + apply blast + using fB apply blast + unfolding d[symmetric] + apply (rule card_image_le) + apply (rule fB) + done + have th0: "dim ?U \ card (f ` B)" + apply (rule span_card_ge_dim) + apply blast + unfolding linear_span_image[OF lf] + apply (rule subset_trans[where B = "f ` UNIV"]) + using sf unfolding surj_def + apply blast + apply (rule image_mono) + apply (rule B(3)) + apply (metis finite_imageI fB) + done + moreover have "card (f ` B) \ card B" + by (rule card_image_le, rule fB) + ultimately have th1: "card B = card (f ` B)" + unfolding d by arith + have fiB: "inj_on f B" + unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] + by blast + from linear_indep_image_lemma[OF lf fB fBi fiB x] fx + have "x = 0" by blast + } + then show ?thesis + unfolding linear_inj_on_iff_eq_0[OF lf subspace_UNIV] + using B(3) + by blast +qed + +lemma linear_inverse_left: + assumes lf: "linear scale scale f" + and lf': "linear scale scale f'" + shows "f \ f' = id \ f' \ f = id" +proof - + { + fix f f':: "'b \ 'b" + assume lf: "linear scale scale f" "linear scale scale f'" + assume f: "f \ f' = id" + from f have sf: "surj f" + apply (auto simp add: o_def id_def surj_def) + apply metis + done + interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales + from linear_surjective_isomorphism[OF lf(1) sf] lf f + have "f' \ f = id" + unfolding fun_eq_iff o_def id_def by metis + } + then show ?thesis + using lf lf' by metis +qed + +lemma left_inverse_linear: + assumes lf: "linear scale scale f" + and gf: "g \ f = id" + shows "linear scale scale g" +proof - + from gf have fi: "inj f" + apply (auto simp add: inj_on_def o_def id_def fun_eq_iff) + apply metis + done + interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales + from linear_injective_isomorphism[OF lf fi] + obtain h :: "'b \ 'b" where h: "linear scale scale h" "\x. h (f x) = x" "\x. f (h x) = x" + by blast + have "h = g" + apply (rule ext) using gf h(2,3) + apply (simp add: o_def id_def fun_eq_iff) + apply metis + done + with h(1) show ?thesis by blast +qed + +lemma inj_linear_imp_inv_linear: + assumes "linear scale scale f" "inj f" shows "linear scale scale (inv f)" + using assms inj_iff left_inverse_linear by blast + +lemma right_inverse_linear: + assumes lf: "linear scale scale f" + and gf: "f \ g = id" + shows "linear scale scale g" +proof - + from gf have fi: "surj f" + by (auto simp add: surj_def o_def id_def) metis + interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales + from linear_surjective_isomorphism[OF lf fi] + obtain h:: "'b \ 'b" where h: "linear scale scale h" "\x. h (f x) = x" "\x. f (h x) = x" + by blast + have "h = g" + apply (rule ext) + using gf h(2,3) + apply (simp add: o_def id_def fun_eq_iff) + apply metis + done + with h(1) show ?thesis by blast +qed + +end + +context finite_dimensional_vector_space_pair begin + +lemma subspace_isomorphism: + assumes s: "vs1.subspace S" + and t: "vs2.subspace T" + and d: "vs1.dim S = vs2.dim T" + shows "\f. linear s1 s2 f \ f ` S = T \ inj_on f S" +proof - + from vs1.basis_exists[of S] vs1.finiteI_independent + obtain B where B: "B \ S" "vs1.independent B" "S \ vs1.span B" "card B = vs1.dim S" and fB: "finite B" + by blast + from vs2.basis_exists[of T] vs2.finiteI_independent + obtain C where C: "C \ T" "vs2.independent C" "T \ vs2.span C" "card C = vs2.dim T" and fC: "finite C" + by blast + from B(4) C(4) card_le_inj[of B C] d + obtain f where f: "f ` B \ C" "inj_on f B" using \finite B\ \finite C\ + by auto + from linear_independent_extend[OF B(2)] + obtain g where g: "linear s1 s2 g" "\x\ B. g x = f x" + by blast + interpret g: linear s1 s2 g by fact + from inj_on_iff_eq_card[OF fB, of f] f(2) have "card (f ` B) = card B" + by simp + with B(4) C(4) have ceq: "card (f ` B) = card C" + using d by simp + have "g ` B = f ` B" + using g(2) by (auto simp add: image_iff) + also have "\ = C" using card_subset_eq[OF fC f(1) ceq] . + finally have gBC: "g ` B = C" . + have gi: "inj_on g B" + using f(2) g(2) by (auto simp add: inj_on_def) + note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi] + { + fix x y + assume x: "x \ S" and y: "y \ S" and gxy: "g x = g y" + from B(3) x y have x': "x \ vs1.span B" and y': "y \ vs1.span B" + by blast+ + from gxy have th0: "g (x - y) = 0" + by (simp add: linear_diff g) + have th1: "x - y \ vs1.span B" + using x' y' by (metis vs1.span_diff) + have "x = y" + using g0[OF th1 th0] by simp + } + then have giS: "inj_on g S" + unfolding inj_on_def by blast + from vs1.span_subspace[OF B(1,3) s] have "g ` S = vs2.span (g ` B)" + by (simp add: module_hom.span_image[OF g(1)[unfolded linear_iff_module_hom]]) + also have "\ = vs2.span C" unfolding gBC .. + also have "\ = T" using vs2.span_subspace[OF C(1,3) t] . + finally have gS: "g ` S = T" . + from g(1) gS giS show ?thesis + by blast +qed + +end + +hide_const (open) linear + +end \ No newline at end of file diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/ex/Ballot.thy --- a/src/HOL/ex/Ballot.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/ex/Ballot.thy Wed May 02 13:49:38 2018 +0200 @@ -8,7 +8,6 @@ theory Ballot imports Complex_Main - "HOL-Library.FuncSet" begin subsection \Preliminaries\ diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/ex/Birthday_Paradox.thy --- a/src/HOL/ex/Birthday_Paradox.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/ex/Birthday_Paradox.thy Wed May 02 13:49:38 2018 +0200 @@ -5,7 +5,7 @@ section \A Formulation of the Birthday Paradox\ theory Birthday_Paradox -imports Main "HOL-Library.FuncSet" +imports Main HOL.FuncSet begin section \Cardinality\ diff -r 0a2a1b6507c1 -r 493b818e8e10 src/HOL/ex/Tarski.thy --- a/src/HOL/ex/Tarski.thy Wed Apr 18 21:12:50 2018 +0100 +++ b/src/HOL/ex/Tarski.thy Wed May 02 13:49:38 2018 +0200 @@ -5,7 +5,7 @@ section \The Full Theorem of Tarski\ theory Tarski -imports Main "HOL-Library.FuncSet" +imports Main HOL.FuncSet begin text \