src/HOL/Vector_Spaces.thy
author immler
Fri, 13 Jul 2018 12:14:26 +0200
changeset 68620 707437105595
parent 68412 07f8c09e3f79
child 68627 e371784becd9
permissions -rw-r--r--
relaxed assumptions for dim_image_eq and dim_image_le
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68189
6163c90694ef tuned headers;
wenzelm
parents: 68188
diff changeset
     1
(*  Title:      HOL/Vector_Spaces.thy
6163c90694ef tuned headers;
wenzelm
parents: 68188
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
6163c90694ef tuned headers;
wenzelm
parents: 68188
diff changeset
     3
    Author:     Jose Divasón <jose.divasonm at unirioja.es>
6163c90694ef tuned headers;
wenzelm
parents: 68188
diff changeset
     4
    Author:     Jesús Aransay <jesus-maria.aransay at unirioja.es>
6163c90694ef tuned headers;
wenzelm
parents: 68188
diff changeset
     5
    Author:     Johannes Hölzl, VU Amsterdam
6163c90694ef tuned headers;
wenzelm
parents: 68188
diff changeset
     6
    Author:     Fabian Immler, TUM
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
     7
*)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
     8
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
     9
section \<open>Vector Spaces\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    10
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    11
theory Vector_Spaces
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
    12
  imports Modules
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    13
begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    14
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    15
lemma isomorphism_expand:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    16
  "f \<circ> g = id \<and> g \<circ> f = id \<longleftrightarrow> (\<forall>x. f (g x) = x) \<and> (\<forall>x. g (f x) = x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    17
  by (simp add: fun_eq_iff o_def id_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    18
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    19
lemma left_right_inverse_eq:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    20
  assumes fg: "f \<circ> g = id"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    21
    and gh: "g \<circ> h = id"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    22
  shows "f = h"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    23
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    24
  have "f = f \<circ> (g \<circ> h)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    25
    unfolding gh by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    26
  also have "\<dots> = (f \<circ> g) \<circ> h"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    27
    by (simp add: o_assoc)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    28
  finally show "f = h"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    29
    unfolding fg by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    30
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    31
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    32
lemma ordLeq3_finite_infinite:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    33
  assumes A: "finite A" and B: "infinite B" shows "ordLeq3 (card_of A) (card_of B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    34
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    35
  have \<open>ordLeq3 (card_of A) (card_of B) \<or> ordLeq3 (card_of B) (card_of A)\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    36
    by (intro ordLeq_total card_of_Well_order)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    37
  moreover have "\<not> ordLeq3 (card_of B) (card_of A)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    38
    using B A card_of_ordLeq_finite[of B A] by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    39
  ultimately show ?thesis by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    40
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    41
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    42
locale vector_space =
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    43
  fixes scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr "*s" 75)
68397
cace81744c61 isabelle update_comments;
wenzelm
parents: 68189
diff changeset
    44
  assumes vector_space_assms:\<comment> \<open>re-stating the assumptions of \<open>module\<close> instead of extending \<open>module\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    45
   allows us to rewrite in the sublocale.\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    46
    "a *s (x + y) = a *s x + a *s y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    47
    "(a + b) *s x = a *s x + b *s x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    48
    "a *s (b *s x) = (a * b) *s x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    49
    "1 *s x = x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    50
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    51
lemma module_iff_vector_space: "module s \<longleftrightarrow> vector_space s"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    52
  unfolding module_def vector_space_def ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    53
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    54
locale linear = vs1: vector_space s1 + vs2: vector_space s2 + module_hom s1 s2 f
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    55
  for s1 :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr "*a" 75)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    56
  and s2 :: "'a::field \<Rightarrow> 'c::ab_group_add \<Rightarrow> 'c" (infixr "*b" 75)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    57
  and f :: "'b \<Rightarrow> 'c"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    58
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    59
lemma module_hom_iff_linear: "module_hom s1 s2 f \<longleftrightarrow> linear s1 s2 f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    60
  unfolding module_hom_def linear_def module_iff_vector_space by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    61
lemmas module_hom_eq_linear = module_hom_iff_linear[abs_def, THEN meta_eq_to_obj_eq]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    62
lemmas linear_iff_module_hom = module_hom_iff_linear[symmetric]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    63
lemmas linear_module_homI = module_hom_iff_linear[THEN iffD1]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    64
  and module_hom_linearI = module_hom_iff_linear[THEN iffD2]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    65
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    66
context vector_space begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    67
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    68
sublocale module scale rewrites "module_hom = linear"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    69
  by (unfold_locales) (fact vector_space_assms module_hom_eq_linear)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    70
68397
cace81744c61 isabelle update_comments;
wenzelm
parents: 68189
diff changeset
    71
lemmas\<comment> \<open>from \<open>module\<close>\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    72
      linear_id = module_hom_id
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    73
  and linear_ident = module_hom_ident
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    74
  and linear_scale_self = module_hom_scale_self
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    75
  and linear_scale_left = module_hom_scale_left
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    76
  and linear_uminus = module_hom_uminus
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    77
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    78
lemma linear_imp_scale:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    79
  fixes D::"'a \<Rightarrow> 'b"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    80
  assumes "linear ( *) scale D"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    81
  obtains d where "D = (\<lambda>x. scale x d)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    82
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    83
  interpret linear "( *)" scale D by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    84
  show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    85
    by (metis mult.commute mult.left_neutral scale that)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    86
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    87
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    88
lemma scale_eq_0_iff [simp]: "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    89
  by (metis scale_left_commute right_inverse scale_one scale_scale scale_zero_left)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    90
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    91
lemma scale_left_imp_eq:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    92
  assumes nonzero: "a \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    93
    and scale: "scale a x = scale a y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    94
  shows "x = y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    95
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    96
  from scale have "scale a (x - y) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    97
     by (simp add: scale_right_diff_distrib)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    98
  with nonzero have "x - y = 0" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    99
  then show "x = y" by (simp only: right_minus_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   100
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   101
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   102
lemma scale_right_imp_eq:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   103
  assumes nonzero: "x \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   104
    and scale: "scale a x = scale b x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   105
  shows "a = b"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   106
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   107
  from scale have "scale (a - b) x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   108
     by (simp add: scale_left_diff_distrib)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   109
  with nonzero have "a - b = 0" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   110
  then show "a = b" by (simp only: right_minus_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   111
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   112
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   113
lemma scale_cancel_left [simp]: "scale a x = scale a y \<longleftrightarrow> x = y \<or> a = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   114
  by (auto intro: scale_left_imp_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   115
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   116
lemma scale_cancel_right [simp]: "scale a x = scale b x \<longleftrightarrow> a = b \<or> x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   117
  by (auto intro: scale_right_imp_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   118
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   119
lemma injective_scale: "c \<noteq> 0 \<Longrightarrow> inj (\<lambda>x. scale c x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   120
  by (simp add: inj_on_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   121
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   122
lemma dependent_def: "dependent P \<longleftrightarrow> (\<exists>a \<in> P. a \<in> span (P - {a}))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   123
  unfolding dependent_explicit
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   124
proof safe
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   125
  fix a assume aP: "a \<in> P" and "a \<in> span (P - {a})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   126
  then obtain a S u
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   127
    where aP: "a \<in> P" and fS: "finite S" and SP: "S \<subseteq> P" "a \<notin> S" and ua: "(\<Sum>v\<in>S. u v *s v) = a"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   128
    unfolding span_explicit by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   129
  let ?S = "insert a S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   130
  let ?u = "\<lambda>y. if y = a then - 1 else u y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   131
  from fS SP have "(\<Sum>v\<in>?S. ?u v *s v) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   132
    by (simp add: if_distrib[of "\<lambda>r. r *s a" for a] sum.If_cases field_simps Diff_eq[symmetric] ua)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   133
  moreover have "finite ?S" "?S \<subseteq> P" "a \<in> ?S" "?u a \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   134
    using fS SP aP by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   135
  ultimately show "\<exists>t u. finite t \<and> t \<subseteq> P \<and> (\<Sum>v\<in>t. u v *s v) = 0 \<and> (\<exists>v\<in>t. u v \<noteq> 0)" by fast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   136
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   137
  fix S u v
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   138
  assume fS: "finite S" and SP: "S \<subseteq> P" and vS: "v \<in> S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   139
   and uv: "u v \<noteq> 0" and u: "(\<Sum>v\<in>S. u v *s v) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   140
  let ?a = v
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   141
  let ?S = "S - {v}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   142
  let ?u = "\<lambda>i. (- u i) / u v"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   143
  have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   144
    using fS SP vS by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   145
  have "(\<Sum>v\<in>?S. ?u v *s v) = (\<Sum>v\<in>S. (- (inverse (u ?a))) *s (u v *s v)) - ?u v *s v"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   146
    using fS vS uv by (simp add: sum_diff1 field_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   147
  also have "\<dots> = ?a"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   148
    unfolding scale_sum_right[symmetric] u using uv by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   149
  finally have "(\<Sum>v\<in>?S. ?u v *s v) = ?a" .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   150
  with th0 show "\<exists>a \<in> P. a \<in> span (P - {a})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   151
    unfolding span_explicit by (auto intro!: bexI[where x="?a"] exI[where x="?S"] exI[where x="?u"])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   152
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   153
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   154
lemma dependent_single[simp]: "dependent {x} \<longleftrightarrow> x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   155
  unfolding dependent_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   156
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   157
lemma in_span_insert:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   158
  assumes a: "a \<in> span (insert b S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   159
    and na: "a \<notin> span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   160
  shows "b \<in> span (insert a S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   161
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   162
  from span_breakdown[of b "insert b S" a, OF insertI1 a]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   163
  obtain k where k: "a - k *s b \<in> span (S - {b})" by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   164
  have "k \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   165
  proof
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   166
    assume "k = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   167
    with k span_mono[of "S - {b}" S] have "a \<in> span S" by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   168
    with na show False by blast  
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   169
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   170
  then have eq: "b = (1/k) *s a - (1/k) *s (a - k *s b)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   171
    by (simp add: algebra_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   172
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   173
  from k have "(1/k) *s (a - k *s b) \<in> span (S - {b})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   174
    by (rule span_scale)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   175
  also have "... \<subseteq> span (insert a S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   176
    by (rule span_mono) auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   177
  finally show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   178
    using k by (subst eq) (blast intro: span_diff span_scale span_base)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   179
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   180
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   181
lemma dependent_insertD: assumes a: "a \<notin> span S" and S: "dependent (insert a S)" shows "dependent S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   182
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   183
  have "a \<notin> S" using a by (auto dest: span_base)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   184
  obtain b where b: "b = a \<or> b \<in> S" "b \<in> span (insert a S - {b})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   185
    using S unfolding dependent_def by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   186
  have "b \<noteq> a" "b \<in> S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   187
    using b \<open>a \<notin> S\<close> a by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   188
  with b have *: "b \<in> span (insert a (S - {b}))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   189
    by (auto simp: insert_Diff_if)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   190
  show "dependent S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   191
  proof cases
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   192
    assume "b \<in> span (S - {b})" with \<open>b \<in> S\<close> show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   193
      by (auto simp add: dependent_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   194
  next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   195
    assume "b \<notin> span (S - {b})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   196
    with * have "a \<in> span (insert b (S - {b}))" by (rule in_span_insert)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   197
    with a show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   198
      using \<open>b \<in> S\<close> by (auto simp: insert_absorb)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   199
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   200
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   201
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   202
lemma independent_insertI: "a \<notin> span S \<Longrightarrow> independent S \<Longrightarrow> independent (insert a S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   203
  by (auto dest: dependent_insertD)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   204
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   205
lemma independent_insert:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   206
  "independent (insert a S) \<longleftrightarrow> (if a \<in> S then independent S else independent S \<and> a \<notin> span S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   207
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   208
  have "a \<notin> S \<Longrightarrow> a \<in> span S \<Longrightarrow> dependent (insert a S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   209
    by (auto simp: dependent_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   210
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   211
    by (auto intro: dependent_mono simp: independent_insertI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   212
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   213
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   214
lemma maximal_independent_subset_extend:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   215
  assumes "S \<subseteq> V" "independent S"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   216
  obtains B where "S \<subseteq> B" "B \<subseteq> V" "independent B" "V \<subseteq> span B"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   217
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   218
  let ?C = "{B. S \<subseteq> B \<and> independent B \<and> B \<subseteq> V}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   219
  have "\<exists>M\<in>?C. \<forall>X\<in>?C. M \<subseteq> X \<longrightarrow> X = M"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   220
  proof (rule subset_Zorn)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   221
    fix C :: "'b set set" assume "subset.chain ?C C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   222
    then have C: "\<And>c. c \<in> C \<Longrightarrow> c \<subseteq> V" "\<And>c. c \<in> C \<Longrightarrow> S \<subseteq> c" "\<And>c. c \<in> C \<Longrightarrow> independent c"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   223
      "\<And>c d. c \<in> C \<Longrightarrow> d \<in> C \<Longrightarrow> c \<subseteq> d \<or> d \<subseteq> c"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   224
      unfolding subset.chain_def by blast+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   225
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   226
    show "\<exists>U\<in>?C. \<forall>X\<in>C. X \<subseteq> U"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   227
    proof cases
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   228
      assume "C = {}" with assms show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   229
        by (auto intro!: exI[of _ S])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   230
    next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   231
      assume "C \<noteq> {}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   232
      with C(2) have "S \<subseteq> \<Union>C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   233
        by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   234
      moreover have "independent (\<Union>C)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   235
        by (intro independent_Union_directed C)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   236
      moreover have "\<Union>C \<subseteq> V"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   237
        using C by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   238
      ultimately show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   239
        by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   240
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   241
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   242
  then obtain B where B: "independent B" "B \<subseteq> V" "S \<subseteq> B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   243
    and max: "\<And>S. independent S \<Longrightarrow> S \<subseteq> V \<Longrightarrow> B \<subseteq> S \<Longrightarrow> S = B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   244
    by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   245
  moreover
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   246
  { assume "\<not> V \<subseteq> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   247
    then obtain v where "v \<in> V" "v \<notin> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   248
      by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   249
    with B have "independent (insert v B)" by (auto intro: dependent_insertD)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   250
    from max[OF this] \<open>v \<in> V\<close> \<open>B \<subseteq> V\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   251
    have "v \<in> B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   252
      by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   253
    with \<open>v \<notin> span B\<close> have False
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   254
      by (auto intro: span_base) }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   255
  ultimately show ?thesis
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   256
    by (meson that)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   257
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   258
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   259
lemma maximal_independent_subset:
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   260
  obtains B where "B \<subseteq> V" "independent B" "V \<subseteq> span B"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   261
  by (metis maximal_independent_subset_extend[of "{}"] empty_subsetI independent_empty)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   262
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   263
text \<open>Extends a basis from B to a basis of the entire space.\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   264
definition extend_basis :: "'b set \<Rightarrow> 'b set"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   265
  where "extend_basis B = (SOME B'. B \<subseteq> B' \<and> independent B' \<and> span B' = UNIV)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   266
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   267
lemma
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   268
  assumes B: "independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   269
  shows extend_basis_superset: "B \<subseteq> extend_basis B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   270
    and independent_extend_basis: "independent (extend_basis B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   271
    and span_extend_basis[simp]: "span (extend_basis B) = UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   272
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   273
  define p where "p B' \<equiv> B \<subseteq> B' \<and> independent B' \<and> span B' = UNIV" for B'
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   274
  obtain B' where "p B'"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   275
    using maximal_independent_subset_extend[OF subset_UNIV B]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   276
    by (metis top.extremum_uniqueI p_def)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   277
  then have "p (extend_basis B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   278
    unfolding extend_basis_def p_def[symmetric] by (rule someI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   279
  then show "B \<subseteq> extend_basis B" "independent (extend_basis B)" "span (extend_basis B) = UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   280
    by (auto simp: p_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   281
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   282
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   283
lemma in_span_delete:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   284
  assumes a: "a \<in> span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   285
    and na: "a \<notin> span (S - {b})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   286
  shows "b \<in> span (insert a (S - {b}))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   287
  apply (rule in_span_insert)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   288
  apply (rule set_rev_mp)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   289
  apply (rule a)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   290
  apply (rule span_mono)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   291
  apply blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   292
  apply (rule na)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   293
  done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   294
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   295
lemma span_redundant: "x \<in> span S \<Longrightarrow> span (insert x S) = span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   296
  unfolding span_def by (rule hull_redundant)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   297
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   298
lemma span_trans: "x \<in> span S \<Longrightarrow> y \<in> span (insert x S) \<Longrightarrow> y \<in> span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   299
  by (simp only: span_redundant)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   300
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   301
lemma span_insert_0[simp]: "span (insert 0 S) = span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   302
  by (metis span_zero span_redundant)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   303
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   304
lemma span_delete_0 [simp]: "span(S - {0}) = span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   305
proof
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   306
  show "span (S - {0}) \<subseteq> span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   307
    by (blast intro!: span_mono)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   308
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   309
  have "span S \<subseteq> span(insert 0 (S - {0}))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   310
    by (blast intro!: span_mono)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   311
  also have "... \<subseteq> span(S - {0})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   312
    using span_insert_0 by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   313
  finally show "span S \<subseteq> span (S - {0})" .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   314
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   315
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   316
lemma span_image_scale:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   317
  assumes "finite S" and nz: "\<And>x. x \<in> S \<Longrightarrow> c x \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   318
    shows "span ((\<lambda>x. c x *s x) ` S) = span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   319
using assms
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   320
proof (induction S arbitrary: c)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   321
  case (empty c) show ?case by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   322
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   323
  case (insert x F c)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   324
  show ?case
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   325
  proof (intro set_eqI iffI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   326
    fix y
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   327
      assume "y \<in> span ((\<lambda>x. c x *s x) ` insert x F)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   328
      then show "y \<in> span (insert x F)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   329
        using insert by (force simp: span_breakdown_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   330
  next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   331
    fix y
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   332
      assume "y \<in> span (insert x F)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   333
      then show "y \<in> span ((\<lambda>x. c x *s x) ` insert x F)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   334
        using insert
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   335
        apply (clarsimp simp: span_breakdown_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   336
        apply (rule_tac x="k / c x" in exI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   337
        by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   338
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   339
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   340
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   341
lemma exchange_lemma:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   342
  assumes f: "finite T"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   343
    and i: "independent S"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   344
    and sp: "S \<subseteq> span T"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   345
  shows "\<exists>t'. card t' = card T \<and> finite t' \<and> S \<subseteq> t' \<and> t' \<subseteq> S \<union> T \<and> S \<subseteq> span t'"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   346
  using f i sp
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   347
proof (induct "card (T - S)" arbitrary: S T rule: less_induct)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   348
  case less
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   349
  note ft = \<open>finite T\<close> and S = \<open>independent S\<close> and sp = \<open>S \<subseteq> span T\<close>
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   350
  let ?P = "\<lambda>t'. card t' = card T \<and> finite t' \<and> S \<subseteq> t' \<and> t' \<subseteq> S \<union> T \<and> S \<subseteq> span t'"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   351
  show ?case
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   352
  proof (cases "S \<subseteq> T \<or> T \<subseteq> S")
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   353
    case True
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   354
    then show ?thesis
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   355
    proof
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   356
      assume "S \<subseteq> T" then show ?thesis
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   357
        by (metis ft Un_commute sp sup_ge1)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   358
    next
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   359
      assume "T \<subseteq> S" then show ?thesis
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   360
        by (metis Un_absorb sp spanning_subset_independent[OF _ S sp] ft)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   361
    qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   362
  next
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   363
    case False
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   364
    then have st: "\<not> S \<subseteq> T" "\<not> T \<subseteq> S"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   365
      by auto
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   366
    from st(2) obtain b where b: "b \<in> T" "b \<notin> S"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   367
      by blast
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   368
    from b have "T - {b} - S \<subset> T - S"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   369
      by blast
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   370
    then have cardlt: "card (T - {b} - S) < card (T - S)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   371
      using ft by (auto intro: psubset_card_mono)
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   372
    from b ft have ct0: "card T \<noteq> 0"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   373
      by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   374
    show ?thesis
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   375
    proof (cases "S \<subseteq> span (T - {b})")
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   376
      case True
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   377
      from ft have ftb: "finite (T - {b})"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   378
        by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   379
      from less(1)[OF cardlt ftb S True]
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   380
      obtain U where U: "card U = card (T - {b})" "S \<subseteq> U" "U \<subseteq> S \<union> (T - {b})" "S \<subseteq> span U"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   381
        and fu: "finite U" by blast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   382
      let ?w = "insert b U"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   383
      have th0: "S \<subseteq> insert b U"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   384
        using U by blast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   385
      have th1: "insert b U \<subseteq> S \<union> T"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   386
        using U b by blast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   387
      have bu: "b \<notin> U"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   388
        using b U by blast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   389
      from U(1) ft b have "card U = (card T - 1)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   390
        by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   391
      then have th2: "card (insert b U) = card T"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   392
        using card_insert_disjoint[OF fu bu] ct0 by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   393
      from U(4) have "S \<subseteq> span U" .
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   394
      also have "\<dots> \<subseteq> span (insert b U)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   395
        by (rule span_mono) blast
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   396
      finally have th3: "S \<subseteq> span (insert b U)" .
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   397
      from th0 th1 th2 th3 fu have th: "?P ?w"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   398
        by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   399
      from th show ?thesis by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   400
    next
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   401
      case False
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   402
      then obtain a where a: "a \<in> S" "a \<notin> span (T - {b})"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   403
        by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   404
      have ab: "a \<noteq> b"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   405
        using a b by blast
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   406
      have at: "a \<notin> T"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   407
        using a ab span_base[of a "T- {b}"] by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   408
      have mlt: "card ((insert a (T - {b})) - S) < card (T - S)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   409
        using cardlt ft a b by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   410
      have ft': "finite (insert a (T - {b}))"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   411
        using ft by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   412
      have sp': "S \<subseteq> span (insert a (T - {b}))"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   413
      proof
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   414
        fix x
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   415
        assume xs: "x \<in> S"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   416
        have T: "T \<subseteq> insert b (insert a (T - {b}))"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   417
          using b by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   418
        have bs: "b \<in> span (insert a (T - {b}))"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   419
          by (rule in_span_delete) (use a sp in auto)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   420
        from xs sp have "x \<in> span T"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   421
          by blast
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   422
        with span_mono[OF T] have x: "x \<in> span (insert b (insert a (T - {b})))" ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   423
        from span_trans[OF bs x] show "x \<in> span (insert a (T - {b}))" .
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   424
      qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   425
      from less(1)[OF mlt ft' S sp'] obtain U where U:
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   426
        "card U = card (insert a (T - {b}))"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   427
        "finite U" "S \<subseteq> U" "U \<subseteq> S \<union> insert a (T - {b})"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   428
        "S \<subseteq> span U" by blast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   429
      from U a b ft at ct0 have "?P U"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   430
        by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   431
      then show ?thesis by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   432
    qed
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   433
  qed
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   434
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   435
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   436
lemma independent_span_bound:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   437
  assumes f: "finite T"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   438
    and i: "independent S"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   439
    and sp: "S \<subseteq> span T"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   440
  shows "finite S \<and> card S \<le> card T"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   441
  by (metis exchange_lemma[OF f i sp] finite_subset card_mono)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   442
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   443
lemma independent_explicit_finite_subsets:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   444
  "independent A \<longleftrightarrow> (\<forall>S \<subseteq> A. finite S \<longrightarrow> (\<forall>u. (\<Sum>v\<in>S. u v *s v) = 0 \<longrightarrow> (\<forall>v\<in>S. u v = 0)))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   445
  unfolding dependent_explicit [of A] by (simp add: disj_not2)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   446
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   447
lemma independent_if_scalars_zero:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   448
  assumes fin_A: "finite A"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   449
  and sum: "\<And>f x. (\<Sum>x\<in>A. f x *s x) = 0 \<Longrightarrow> x \<in> A \<Longrightarrow> f x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   450
  shows "independent A"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   451
proof (unfold independent_explicit_finite_subsets, clarify)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   452
  fix S v and u :: "'b \<Rightarrow> 'a"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   453
  assume S: "S \<subseteq> A" and v: "v \<in> S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   454
  let ?g = "\<lambda>x. if x \<in> S then u x else 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   455
  have "(\<Sum>v\<in>A. ?g v *s v) = (\<Sum>v\<in>S. u v *s v)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   456
    using S fin_A by (auto intro!: sum.mono_neutral_cong_right)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   457
  also assume "(\<Sum>v\<in>S. u v *s v) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   458
  finally have "?g v = 0" using v S sum by force
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   459
  thus "u v = 0"  unfolding if_P[OF v] .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   460
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   461
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   462
lemma bij_if_span_eq_span_bases:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   463
  assumes B: "independent B" and C: "independent C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   464
    and eq: "span B = span C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   465
  shows "\<exists>f. bij_betw f B C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   466
proof cases
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   467
  assume "finite B \<or> finite C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   468
  then have "finite B \<and> finite C \<and> card C = card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   469
    using independent_span_bound[of B C] independent_span_bound[of C B] assms
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   470
      span_superset[of B] span_superset[of C]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   471
    by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   472
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   473
    by (auto intro!: finite_same_card_bij)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   474
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   475
  assume "\<not> (finite B \<or> finite C)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   476
  then have "infinite B" "infinite C" by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   477
  { fix B C assume  B: "independent B" and C: "independent C" and "infinite B" "infinite C" and eq: "span B = span C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   478
    let ?R = "representation B" and ?R' = "representation C" let ?U = "\<lambda>c. {v. ?R c v \<noteq> 0}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   479
    have in_span_C [simp, intro]: \<open>b \<in> B \<Longrightarrow> b \<in> span C\<close> for b unfolding eq[symmetric] by (rule span_base) 
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   480
    have in_span_B [simp, intro]: \<open>c \<in> C \<Longrightarrow> c \<in> span B\<close> for c unfolding eq by (rule span_base) 
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   481
    have \<open>B \<subseteq> (\<Union>c\<in>C. ?U c)\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   482
    proof
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   483
      fix b assume \<open>b \<in> B\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   484
      have \<open>b \<in> span C\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   485
        using \<open>b \<in> B\<close> unfolding eq[symmetric] by (rule span_base)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   486
      have \<open>(\<Sum>v | ?R' b v \<noteq> 0. \<Sum>w | ?R v w \<noteq> 0. (?R' b v * ?R v w) *s w) =
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   487
          (\<Sum>v | ?R' b v \<noteq> 0. ?R' b v *s (\<Sum>w | ?R v w \<noteq> 0. ?R v w *s w))\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   488
        by (simp add: scale_sum_right)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   489
      also have \<open>\<dots> = (\<Sum>v | ?R' b v \<noteq> 0. ?R' b v *s v)\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   490
        by (auto simp: sum_nonzero_representation_eq B eq span_base representation_ne_zero)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   491
      also have \<open>\<dots> = b\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   492
        by (rule sum_nonzero_representation_eq[OF C \<open>b \<in> span C\<close>])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   493
      finally have "?R b b = ?R (\<Sum>v | ?R' b v \<noteq> 0. \<Sum>w | ?R v w \<noteq> 0. (?R' b v * ?R v w) *s w) b"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   494
        by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   495
      also have \<open>\<dots> = (\<Sum>v | ?R' b v \<noteq> 0. \<Sum>w | ?R v w \<noteq> 0. ?R' b v * ?R v w * ?R w b)\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   496
        using B \<open>b \<in> B\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   497
        apply (subst representation_sum[OF B])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   498
         apply (fastforce intro: span_sum span_scale span_base representation_ne_zero)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   499
        apply (rule sum.cong[OF refl])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   500
        apply (subst representation_sum[OF B])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   501
         apply (simp add: span_sum span_scale span_base representation_ne_zero)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   502
        apply (simp add: representation_scale[OF B] span_base representation_ne_zero)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   503
        done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   504
      finally have "(\<Sum>v | ?R' b v \<noteq> 0. \<Sum>w | ?R v w \<noteq> 0. ?R' b v * ?R v w * ?R w b) \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   505
        using representation_basis[OF B \<open>b \<in> B\<close>] by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   506
      then obtain v w where bv: "?R' b v \<noteq> 0" and vw: "?R v w \<noteq> 0" and "?R' b v * ?R v w * ?R w b \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   507
        by (blast elim: sum.not_neutral_contains_not_neutral)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   508
      with representation_basis[OF B, of w] vw[THEN representation_ne_zero]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   509
      have \<open>?R' b v \<noteq> 0\<close> \<open>?R v b \<noteq> 0\<close> by (auto split: if_splits)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   510
      then show \<open>b \<in> (\<Union>c\<in>C. ?U c)\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   511
        by (auto dest: representation_ne_zero)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   512
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   513
    then have B_eq: \<open>B = (\<Union>c\<in>C. ?U c)\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   514
      by (auto intro: span_base representation_ne_zero eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   515
    have "ordLeq3 (card_of B) (card_of C)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   516
    proof (subst B_eq, rule card_of_UNION_ordLeq_infinite[OF \<open>infinite C\<close>])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   517
      show "ordLeq3 (card_of C) (card_of C)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   518
        by (intro ordLeq_refl card_of_Card_order)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   519
      show "\<forall>c\<in>C. ordLeq3 (card_of {v. representation B c v \<noteq> 0}) (card_of C)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   520
        by (intro ballI ordLeq3_finite_infinite \<open>infinite C\<close> finite_representation)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   521
    qed }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   522
  from this[of B C] this[of C B] B C eq \<open>infinite C\<close> \<open>infinite B\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   523
  show ?thesis by (auto simp add: ordIso_iff_ordLeq card_of_ordIso)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   524
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   525
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   526
definition dim :: "'b set \<Rightarrow> nat"
68412
07f8c09e3f79 default value for parametricity of dim
immler
parents: 68397
diff changeset
   527
  where "dim V = (if \<exists>b. independent b \<and> span b = span V then
07f8c09e3f79 default value for parametricity of dim
immler
parents: 68397
diff changeset
   528
    card (SOME b. independent b \<and> span b = span V) else 0)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   529
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   530
lemma dim_eq_card:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   531
  assumes BV: "span B = span V" and B: "independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   532
  shows "dim V = card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   533
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   534
  define p where "p b \<equiv> independent b \<and> span b = span V" for b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   535
  have "p (SOME B. p B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   536
    using assms by (intro someI[of p B]) (auto simp: p_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   537
  then have "\<exists>f. bij_betw f B (SOME B. p B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   538
    by (subst (asm) p_def, intro bij_if_span_eq_span_bases[OF B]) (simp_all add: BV)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   539
  then have "card B = card (SOME B. p B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   540
    by (auto intro: bij_betw_same_card)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   541
  then show ?thesis
68412
07f8c09e3f79 default value for parametricity of dim
immler
parents: 68397
diff changeset
   542
    using BV B
07f8c09e3f79 default value for parametricity of dim
immler
parents: 68397
diff changeset
   543
    by (auto simp add: dim_def p_def)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   544
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   545
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   546
lemma basis_card_eq_dim: "B \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = dim V"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   547
  using dim_eq_card[of B V] span_mono[of B V] span_minimal[OF _ subspace_span, of V B] by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   548
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   549
lemma basis_exists:
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   550
  obtains B where "B \<subseteq> V" "independent B" "V \<subseteq> span B" "card B = dim V"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   551
  by (meson basis_card_eq_dim empty_subsetI independent_empty maximal_independent_subset_extend)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   552
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   553
lemma dim_eq_card_independent: "independent B \<Longrightarrow> dim B = card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   554
  by (rule dim_eq_card[OF refl])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   555
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   556
lemma dim_span[simp]: "dim (span S) = dim S"
68412
07f8c09e3f79 default value for parametricity of dim
immler
parents: 68397
diff changeset
   557
  by (auto simp add: dim_def span_span)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   558
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   559
lemma dim_span_eq_card_independent: "independent B \<Longrightarrow> dim (span B) = card B"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   560
  by (simp add: dim_eq_card)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   561
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   562
lemma dim_le_card: assumes "V \<subseteq> span W" "finite W" shows "dim V \<le> card W"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   563
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   564
  obtain A where "independent A" "A \<subseteq> V" "V \<subseteq> span A"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   565
    using maximal_independent_subset[of V] by force
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   566
  with assms independent_span_bound[of W A] basis_card_eq_dim[of A V]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   567
  show ?thesis by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   568
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   569
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   570
lemma span_eq_dim: "span S = span T \<Longrightarrow> dim S = dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   571
  by (metis dim_span)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   572
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   573
corollary dim_le_card':
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   574
  "finite s \<Longrightarrow> dim s \<le> card s"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   575
  by (metis basis_exists card_mono)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   576
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   577
lemma span_card_ge_dim:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   578
  "B \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> finite B \<Longrightarrow> dim V \<le> card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   579
  by (simp add: dim_le_card)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   580
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   581
lemma dim_unique:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   582
  "B \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = n \<Longrightarrow> dim V = n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   583
  by (metis basis_card_eq_dim)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   584
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   585
lemma subspace_sums: "\<lbrakk>subspace S; subspace T\<rbrakk> \<Longrightarrow> subspace {x + y|x y. x \<in> S \<and> y \<in> T}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   586
  apply (simp add: subspace_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   587
  apply (intro conjI impI allI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   588
  using add.right_neutral apply blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   589
   apply clarify
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   590
   apply (metis add.assoc add.left_commute)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   591
  using scale_right_distrib by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   592
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   593
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   594
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   595
lemma linear_iff: "linear s1 s2 f \<longleftrightarrow>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   596
  (vector_space s1 \<and> vector_space s2 \<and> (\<forall>x y. f (x + y) = f x + f y) \<and> (\<forall>c x. f (s1 c x) = s2 c (f x)))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   597
  unfolding linear_def module_hom_iff vector_space_def module_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   598
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   599
context begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   600
qualified lemma linear_compose: "linear s1 s2 f \<Longrightarrow> linear s2 s3 g \<Longrightarrow> linear s1 s3 (g o f)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   601
  unfolding module_hom_iff_linear[symmetric]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   602
  by (rule module_hom_compose)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   603
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   604
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   605
locale vector_space_pair = vs1: vector_space s1 + vs2: vector_space s2
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   606
  for s1 :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr "*a" 75)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   607
  and s2 :: "'a::field \<Rightarrow> 'c::ab_group_add \<Rightarrow> 'c" (infixr "*b" 75)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   608
begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   609
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   610
context fixes f assumes "linear s1 s2 f" begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   611
interpretation linear s1 s2 f by fact
68397
cace81744c61 isabelle update_comments;
wenzelm
parents: 68189
diff changeset
   612
lemmas\<comment> \<open>from locale \<open>module_hom\<close>\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   613
      linear_0 = zero
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   614
  and linear_add = add
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   615
  and linear_scale = scale
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   616
  and linear_neg = neg
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   617
  and linear_diff = diff
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   618
  and linear_sum = sum
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   619
  and linear_inj_on_iff_eq_0 = inj_on_iff_eq_0
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   620
  and linear_inj_iff_eq_0 = inj_iff_eq_0
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   621
  and linear_subspace_image = subspace_image
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   622
  and linear_subspace_vimage = subspace_vimage
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   623
  and linear_subspace_kernel = subspace_kernel
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   624
  and linear_span_image = span_image
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   625
  and linear_dependent_inj_imageD = dependent_inj_imageD
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   626
  and linear_eq_0_on_span = eq_0_on_span
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   627
  and linear_independent_injective_image = independent_injective_image
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   628
  and linear_inj_on_span_independent_image = inj_on_span_independent_image
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   629
  and linear_inj_on_span_iff_independent_image = inj_on_span_iff_independent_image
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   630
  and linear_subspace_linear_preimage = subspace_linear_preimage
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   631
  and linear_spans_image = spans_image
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   632
  and linear_spanning_surjective_image = spanning_surjective_image
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   633
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   634
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   635
sublocale module_pair
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   636
  rewrites "module_hom = linear"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   637
  by unfold_locales (fact module_hom_eq_linear)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   638
68397
cace81744c61 isabelle update_comments;
wenzelm
parents: 68189
diff changeset
   639
lemmas\<comment> \<open>from locale \<open>module_pair\<close>\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   640
      linear_eq_on_span = module_hom_eq_on_span
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   641
  and linear_compose_scale_right = module_hom_scale
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   642
  and linear_compose_add = module_hom_add
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   643
  and linear_zero = module_hom_zero
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   644
  and linear_compose_sub = module_hom_sub
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   645
  and linear_compose_neg = module_hom_neg
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   646
  and linear_compose_scale = module_hom_compose_scale
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   647
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   648
lemma linear_indep_image_lemma:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   649
  assumes lf: "linear s1 s2 f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   650
    and fB: "finite B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   651
    and ifB: "vs2.independent (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   652
    and fi: "inj_on f B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   653
    and xsB: "x \<in> vs1.span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   654
    and fx: "f x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   655
  shows "x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   656
  using fB ifB fi xsB fx
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   657
proof (induction B arbitrary: x rule: finite_induct)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   658
  case empty
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   659
  then show ?case by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   660
next
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   661
  case (insert a b x)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   662
  have th0: "f ` b \<subseteq> f ` (insert a b)"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   663
    by (simp add: subset_insertI)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   664
  have ifb: "vs2.independent (f ` b)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   665
    using vs2.independent_mono insert.prems(1) th0 by blast
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   666
  have fib: "inj_on f b"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   667
    using insert.prems(2) by blast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   668
  from vs1.span_breakdown[of a "insert a b", simplified, OF insert.prems(3)]
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   669
  obtain k where k: "x - k *a a \<in> vs1.span (b - {a})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   670
    by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   671
  have "f (x - k *a a) \<in> vs2.span (f ` b)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   672
    unfolding linear_span_image[OF lf]
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   673
    using "insert.hyps"(2) k by auto
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   674
  then have "f x - k *b f a \<in> vs2.span (f ` b)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   675
    by (simp add: linear_diff linear_scale lf)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   676
  then have th: "-k *b f a \<in> vs2.span (f ` b)"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   677
    using insert.prems(4) by simp
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   678
  have xsb: "x \<in> vs1.span b"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   679
  proof (cases "k = 0")
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   680
    case True
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   681
    with k have "x \<in> vs1.span (b - {a})" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   682
    then show ?thesis using vs1.span_mono[of "b - {a}" b]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   683
      by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   684
  next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   685
    case False
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   686
    from inj_on_image_set_diff[OF insert.prems(2), of "insert a b " "{a}", symmetric]
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   687
    have "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   688
    then have "f a \<notin> vs2.span (f ` b)" 
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   689
      using vs2.dependent_def insert.hyps(2) insert.prems(1) by fastforce
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   690
    moreover have "f a \<in> vs2.span (f ` b)"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   691
      using False vs2.span_scale[OF th, of "- 1/ k"] by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   692
    ultimately have False
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   693
      by blast
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   694
    then show ?thesis by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   695
  qed
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   696
  show "x = 0"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   697
    using ifb fib xsb insert.IH insert.prems(4) by blast
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   698
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   699
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   700
lemma linear_eq_on:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   701
  assumes l: "linear s1 s2 f" "linear s1 s2 g"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   702
  assumes x: "x \<in> vs1.span B" and eq: "\<And>b. b \<in> B \<Longrightarrow> f b = g b"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   703
  shows "f x = g x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   704
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   705
  interpret d: linear s1 s2 "\<lambda>x. f x - g x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   706
    using l by (intro linear_compose_sub) (auto simp: module_hom_iff_linear)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   707
  have "f x - g x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   708
    by (rule d.eq_0_on_span[OF _ x]) (auto simp: eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   709
  then show ?thesis by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   710
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   711
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   712
definition construct :: "'b set \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   713
  where "construct B g v = (\<Sum>b | vs1.representation (vs1.extend_basis B) v b \<noteq> 0.
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   714
      vs1.representation (vs1.extend_basis B) v b *b (if b \<in> B then g b else 0))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   715
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   716
lemma construct_cong: "(\<And>b. b \<in> B \<Longrightarrow> f b = g b) \<Longrightarrow> construct B f = construct B g"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   717
  unfolding construct_def by (rule ext, auto intro!: sum.cong)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   718
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   719
lemma linear_construct:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   720
  assumes B[simp]: "vs1.independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   721
  shows "linear s1 s2 (construct B f)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   722
  unfolding module_hom_iff_linear linear_iff
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   723
proof safe
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   724
  have eB[simp]: "vs1.independent (vs1.extend_basis B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   725
    using vs1.independent_extend_basis[OF B] .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   726
  let ?R = "vs1.representation (vs1.extend_basis B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   727
  fix c x y
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   728
  have "construct B f (x + y) =
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   729
    (\<Sum>b\<in>{b. ?R x b \<noteq> 0} \<union> {b. ?R y b \<noteq> 0}. ?R (x + y) b *b (if b \<in> B then f b else 0))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   730
    by (auto intro!: sum.mono_neutral_cong_left simp: vs1.finite_representation vs1.representation_add construct_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   731
  also have "\<dots> = construct B f x + construct B f y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   732
    by (auto simp: construct_def vs1.representation_add vs2.scale_left_distrib sum.distrib
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   733
      intro!: arg_cong2[where f="(+)"] sum.mono_neutral_cong_right vs1.finite_representation)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   734
  finally show "construct B f (x + y) = construct B f x + construct B f y" .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   735
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   736
  show "construct B f (c *a x) = c *b construct B f x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   737
    by (auto simp del: vs2.scale_scale intro!: sum.mono_neutral_cong_left vs1.finite_representation
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   738
      simp add: construct_def vs2.scale_scale[symmetric] vs1.representation_scale vs2.scale_sum_right)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   739
qed intro_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   740
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   741
lemma construct_basis:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   742
  assumes B[simp]: "vs1.independent B" and b: "b \<in> B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   743
  shows "construct B f b = f b"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   744
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   745
  have *: "vs1.representation (vs1.extend_basis B) b = (\<lambda>v. if v = b then 1 else 0)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   746
    using vs1.extend_basis_superset[OF B] b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   747
    by (intro vs1.representation_basis vs1.independent_extend_basis) auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   748
  then have "{v. vs1.representation (vs1.extend_basis B) b v \<noteq> 0} = {b}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   749
    by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   750
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   751
    unfolding construct_def by (simp add: * b)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   752
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   753
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   754
lemma construct_outside:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   755
  assumes B: "vs1.independent B" and v: "v \<in> vs1.span (vs1.extend_basis B - B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   756
  shows "construct B f v = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   757
  unfolding construct_def
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   758
proof (clarsimp intro!: sum.neutral simp del: vs2.scale_eq_0_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   759
  fix b assume "b \<in> B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   760
  then have "vs1.representation (vs1.extend_basis B - B) v b = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   761
    using vs1.representation_ne_zero[of "vs1.extend_basis B - B" v b] by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   762
  moreover have "vs1.representation (vs1.extend_basis B) v = vs1.representation (vs1.extend_basis B - B) v"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   763
    using vs1.representation_extend[OF vs1.independent_extend_basis[OF B] v] by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   764
  ultimately show "vs1.representation (vs1.extend_basis B) v b *b f b = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   765
    by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   766
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   767
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   768
lemma construct_add:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   769
  assumes B[simp]: "vs1.independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   770
  shows "construct B (\<lambda>x. f x + g x) v = construct B f v + construct B g v"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   771
proof (rule linear_eq_on)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   772
  show "v \<in> vs1.span (vs1.extend_basis B)" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   773
  show "b \<in> vs1.extend_basis B \<Longrightarrow> construct B (\<lambda>x. f x + g x) b = construct B f b + construct B g b" for b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   774
    using construct_outside[OF B vs1.span_base, of b] by (cases "b \<in> B") (auto simp: construct_basis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   775
qed (intro linear_compose_add linear_construct B)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   776
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   777
lemma construct_scale:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   778
  assumes B[simp]: "vs1.independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   779
  shows "construct B (\<lambda>x. c *b f x) v = c *b construct B f v"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   780
proof (rule linear_eq_on)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   781
  show "v \<in> vs1.span (vs1.extend_basis B)" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   782
  show "b \<in> vs1.extend_basis B \<Longrightarrow> construct B (\<lambda>x. c *b f x) b = c *b construct B f b" for b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   783
    using construct_outside[OF B vs1.span_base, of b] by (cases "b \<in> B") (auto simp: construct_basis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   784
qed (intro linear_construct module_hom_scale B)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   785
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   786
lemma construct_in_span:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   787
  assumes B[simp]: "vs1.independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   788
  shows "construct B f v \<in> vs2.span (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   789
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   790
  interpret c: linear s1 s2 "construct B f" by (rule linear_construct) fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   791
  let ?R = "vs1.representation B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   792
  have "v \<in> vs1.span ((vs1.extend_basis B - B) \<union> B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   793
    by (auto simp: Un_absorb2 vs1.extend_basis_superset)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   794
  then obtain x y where "v = x + y" "x \<in> vs1.span (vs1.extend_basis B - B)" "y \<in> vs1.span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   795
    unfolding vs1.span_Un by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   796
  moreover have "construct B f (\<Sum>b | ?R y b \<noteq> 0. ?R y b *a b) \<in> vs2.span (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   797
    by (auto simp add: c.sum c.scale construct_basis vs1.representation_ne_zero
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   798
      intro!: vs2.span_sum vs2.span_scale intro: vs2.span_base )
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   799
  ultimately show "construct B f v \<in> vs2.span (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   800
    by (auto simp add: c.add construct_outside vs1.sum_nonzero_representation_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   801
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   802
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   803
lemma linear_compose_sum:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   804
  assumes lS: "\<forall>a \<in> S. linear s1 s2 (f a)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   805
  shows "linear s1 s2 (\<lambda>x. sum (\<lambda>a. f a x) S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   806
proof (cases "finite S")
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   807
  case True
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   808
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   809
    using lS by induct (simp_all add: linear_zero linear_compose_add)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   810
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   811
  case False
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   812
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   813
    by (simp add: linear_zero)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   814
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   815
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   816
lemma in_span_in_range_construct:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   817
  "x \<in> range (construct B f)" if i: "vs1.independent B" and x: "x \<in> vs2.span (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   818
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   819
  interpret linear "( *a)" "( *b)" "construct B f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   820
    using i by (rule linear_construct)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   821
  obtain bb :: "('b \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'b set \<Rightarrow> 'b" where
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   822
    "\<forall>x0 x1 x2. (\<exists>v4. v4 \<in> x2 \<and> x1 v4 \<noteq> x0 v4) = (bb x0 x1 x2 \<in> x2 \<and> x1 (bb x0 x1 x2) \<noteq> x0 (bb x0 x1 x2))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   823
    by moura
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   824
  then have f2: "\<forall>B Ba f fa. (B \<noteq> Ba \<or> bb fa f Ba \<in> Ba \<and> f (bb fa f Ba) \<noteq> fa (bb fa f Ba)) \<or> f ` B = fa ` Ba"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   825
    by (meson image_cong)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   826
  have "vs1.span B \<subseteq> vs1.span (vs1.extend_basis B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   827
    by (simp add: vs1.extend_basis_superset[OF i] vs1.span_mono)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   828
  then show "x \<in> range (construct B f)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   829
    using f2 x by (metis (no_types) construct_basis[OF i, of _ f]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   830
        vs1.span_extend_basis[OF i] set_mp span_image spans_image)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   831
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   832
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   833
lemma range_construct_eq_span:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   834
  "range (construct B f) = vs2.span (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   835
  if "vs1.independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   836
  by (auto simp: that construct_in_span in_span_in_range_construct)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   837
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   838
lemma linear_independent_extend_subspace:
68397
cace81744c61 isabelle update_comments;
wenzelm
parents: 68189
diff changeset
   839
  \<comment> \<open>legacy: use @{term construct} instead\<close>
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   840
  assumes "vs1.independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   841
  shows "\<exists>g. linear s1 s2 g \<and> (\<forall>x\<in>B. g x = f x) \<and> range g = vs2.span (f`B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   842
  by (rule exI[where x="construct B f"])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   843
    (auto simp: linear_construct assms construct_basis range_construct_eq_span)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   844
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   845
lemma linear_independent_extend:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   846
  "vs1.independent B \<Longrightarrow> \<exists>g. linear s1 s2 g \<and> (\<forall>x\<in>B. g x = f x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   847
  using linear_independent_extend_subspace[of B f] by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   848
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   849
lemma linear_exists_left_inverse_on:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   850
  assumes lf: "linear s1 s2 f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   851
  assumes V: "vs1.subspace V" and f: "inj_on f V"
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
   852
  shows "\<exists>g. g ` UNIV \<subseteq> V \<and> linear s2 s1 g \<and> (\<forall>v\<in>V. g (f v) = v)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   853
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   854
  interpret linear s1 s2 f by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   855
  obtain B where V_eq: "V = vs1.span B" and B: "vs1.independent B"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   856
    using vs1.maximal_independent_subset[of V] vs1.span_minimal[OF _ \<open>vs1.subspace V\<close>]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   857
    by (metis antisym_conv)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   858
  have f: "inj_on f (vs1.span B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   859
    using f unfolding V_eq .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   860
  show ?thesis
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
   861
  proof (intro exI ballI conjI)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   862
    interpret p: vector_space_pair s2 s1 by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   863
    have fB: "vs2.independent (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   864
      using independent_injective_image[OF B f] .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   865
    let ?g = "p.construct (f ` B) (the_inv_into B f)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   866
    show "linear ( *b) ( *a) ?g"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   867
      by (rule p.linear_construct[OF fB])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   868
    have "?g b \<in> vs1.span (the_inv_into B f ` f ` B)" for b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   869
      by (intro p.construct_in_span fB)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   870
    moreover have "the_inv_into B f ` f ` B = B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   871
      by (auto simp: image_comp comp_def the_inv_into_f_f inj_on_subset[OF f vs1.span_superset]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   872
          cong: image_cong)
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
   873
    ultimately show "?g ` UNIV \<subseteq> V"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   874
      by (auto simp: V_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   875
    have "(?g \<circ> f) v = id v" if "v \<in> vs1.span B" for v
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   876
    proof (rule vector_space_pair.linear_eq_on[where x=v])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   877
      show "vector_space_pair ( *a) ( *a)" by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   878
      show "linear ( *a) ( *a) (?g \<circ> f)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   879
        apply (rule Vector_Spaces.linear_compose[of _ "( *b)"])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   880
        subgoal by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   881
        apply fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   882
        done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   883
      show "linear ( *a) ( *a) id" by (rule vs1.linear_id)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   884
      show "v \<in> vs1.span B" by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   885
      show "b \<in> B \<Longrightarrow> (p.construct (f ` B) (the_inv_into B f) \<circ> f) b = id b" for b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   886
        by (simp add: p.construct_basis fB the_inv_into_f_f inj_on_subset[OF f vs1.span_superset])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   887
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   888
    then show "v \<in> V \<Longrightarrow> ?g (f v) = v" for v by (auto simp: comp_def id_def V_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   889
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   890
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   891
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   892
lemma linear_exists_right_inverse_on:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   893
  assumes lf: "linear s1 s2 f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   894
  assumes "vs1.subspace V"
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
   895
  shows "\<exists>g. g ` UNIV \<subseteq> V \<and> linear s2 s1 g \<and> (\<forall>v\<in>f ` V. f (g v) = v)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   896
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   897
  obtain B where V_eq: "V = vs1.span B" and B: "vs1.independent B"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   898
    using vs1.maximal_independent_subset[of V] vs1.span_minimal[OF _ \<open>vs1.subspace V\<close>]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   899
    by (metis antisym_conv)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   900
  obtain C where C: "vs2.independent C" and fB_C: "f ` B \<subseteq> vs2.span C" "C \<subseteq> f ` B"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   901
    using vs2.maximal_independent_subset[of "f ` B"] by metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   902
  then have "\<forall>v\<in>C. \<exists>b\<in>B. v = f b" by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   903
  then obtain g where g: "\<And>v. v \<in> C \<Longrightarrow> g v \<in> B" "\<And>v. v \<in> C \<Longrightarrow> f (g v) = v" by metis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   904
  show ?thesis
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
   905
  proof (intro exI ballI conjI)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   906
    interpret p: vector_space_pair s2 s1 by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   907
    let ?g = "p.construct C g"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   908
    show "linear ( *b) ( *a) ?g"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   909
      by (rule p.linear_construct[OF C])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   910
    have "?g v \<in> vs1.span (g ` C)" for v
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   911
      by (rule p.construct_in_span[OF C])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   912
    also have "\<dots> \<subseteq> V" unfolding V_eq using g by (intro vs1.span_mono) auto
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
   913
    finally show "?g ` UNIV \<subseteq> V" by auto
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   914
    have "(f \<circ> ?g) v = id v" if v: "v \<in> f ` V" for v
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   915
    proof (rule vector_space_pair.linear_eq_on[where x=v])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   916
      show "vector_space_pair ( *b) ( *b)" by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   917
      show "linear ( *b) ( *b) (f \<circ> ?g)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   918
        apply (rule Vector_Spaces.linear_compose[of _ "( *a)"])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   919
        apply fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   920
        subgoal by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   921
        done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   922
      show "linear ( *b) ( *b) id" by (rule vs2.linear_id)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   923
      have "vs2.span (f ` B) = vs2.span C"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   924
        using fB_C vs2.span_mono[of C "f ` B"] vs2.span_minimal[of "f`B" "vs2.span C"]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   925
        by auto
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   926
      then show "v \<in> vs2.span C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   927
        using v linear_span_image[OF lf, of B] by (simp add: V_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   928
      show "(f \<circ> p.construct C g) b = id b" if b: "b \<in> C" for b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   929
        by (auto simp: p.construct_basis g C b)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   930
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   931
    then show "v \<in> f ` V \<Longrightarrow> f (?g v) = v" for v by (auto simp: comp_def id_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   932
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   933
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   934
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   935
lemma linear_inj_on_left_inverse:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   936
  assumes lf: "linear s1 s2 f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   937
  assumes fi: "inj_on f (vs1.span S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   938
  shows "\<exists>g. range g \<subseteq> vs1.span S \<and> linear s2 s1 g \<and> (\<forall>x\<in>vs1.span S. g (f x) = x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   939
  using linear_exists_left_inverse_on[OF lf vs1.subspace_span fi]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   940
  by (auto simp: linear_iff_module_hom)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   941
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   942
lemma linear_injective_left_inverse: "linear s1 s2 f \<Longrightarrow> inj f \<Longrightarrow> \<exists>g. linear s2 s1 g \<and> g \<circ> f = id"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   943
  using linear_inj_on_left_inverse[of f UNIV]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   944
  by force
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   945
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   946
lemma linear_surj_right_inverse:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   947
  assumes lf: "linear s1 s2 f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   948
  assumes sf: "vs2.span T \<subseteq> f`vs1.span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   949
  shows "\<exists>g. range g \<subseteq> vs1.span S \<and> linear s2 s1 g \<and> (\<forall>x\<in>vs2.span T. f (g x) = x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   950
  using linear_exists_right_inverse_on[OF lf vs1.subspace_span, of S] sf
68188
2af1f142f855 move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents: 68074
diff changeset
   951
  by (force simp: linear_iff_module_hom)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   952
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   953
lemma linear_surjective_right_inverse: "linear s1 s2 f \<Longrightarrow> surj f \<Longrightarrow> \<exists>g. linear s2 s1 g \<and> f \<circ> g = id"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   954
  using linear_surj_right_inverse[of f UNIV UNIV]
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   955
  by (auto simp: fun_eq_iff)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   956
68620
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   957
lemma finite_basis_to_basis_subspace_isomorphism:
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   958
  assumes s: "vs1.subspace S"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   959
    and t: "vs2.subspace T"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   960
    and d: "vs1.dim S = vs2.dim T"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   961
    and fB: "finite B"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   962
    and B: "B \<subseteq> S" "vs1.independent B" "S \<subseteq> vs1.span B" "card B = vs1.dim S"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   963
    and fC: "finite C"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   964
    and C: "C \<subseteq> T" "vs2.independent C" "T \<subseteq> vs2.span C" "card C = vs2.dim T"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   965
  shows "\<exists>f. linear s1 s2 f \<and> f ` B = C \<and> f ` S = T \<and> inj_on f S"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   966
proof -
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   967
  from B(4) C(4) card_le_inj[of B C] d obtain f where
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   968
    f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close> by auto
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   969
  from linear_independent_extend[OF B(2)] obtain g where
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   970
    g: "linear s1 s2 g" "\<forall>x \<in> B. g x = f x" by blast
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   971
  interpret g: linear s1 s2 g by fact
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   972
  from inj_on_iff_eq_card[OF fB, of f] f(2)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   973
  have "card (f ` B) = card B" by simp
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   974
  with B(4) C(4) have ceq: "card (f ` B) = card C" using d
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   975
    by simp
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   976
  have "g ` B = f ` B" using g(2)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   977
    by (auto simp add: image_iff)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   978
  also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] .
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   979
  finally have gBC: "g ` B = C" .
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   980
  have gi: "inj_on g B" using f(2) g(2)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   981
    by (auto simp add: inj_on_def)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   982
  note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi]
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   983
  {
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   984
    fix x y
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   985
    assume x: "x \<in> S" and y: "y \<in> S" and gxy: "g x = g y"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   986
    from B(3) x y have x': "x \<in> vs1.span B" and y': "y \<in> vs1.span B"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   987
      by blast+
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   988
    from gxy have th0: "g (x - y) = 0"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   989
      by (simp add: g.diff)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   990
    have th1: "x - y \<in> vs1.span B" using x' y'
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   991
      by (metis vs1.span_diff)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   992
    have "x = y" using g0[OF th1 th0] by simp
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   993
  }
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   994
  then have giS: "inj_on g S" unfolding inj_on_def by blast
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   995
  from vs1.span_subspace[OF B(1,3) s]
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   996
  have "g ` S = vs2.span (g ` B)"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   997
    by (simp add: g.span_image)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   998
  also have "\<dots> = vs2.span C"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
   999
    unfolding gBC ..
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1000
  also have "\<dots> = T"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1001
    using vs2.span_subspace[OF C(1,3) t] .
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1002
  finally have gS: "g ` S = T" .
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1003
  from g(1) gS giS gBC show ?thesis
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1004
    by blast
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1005
qed
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1006
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1007
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1008
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1009
lemma surjective_iff_injective_gen:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1010
  assumes fS: "finite S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1011
    and fT: "finite T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1012
    and c: "card S = card T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1013
    and ST: "f ` S \<subseteq> T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1014
  shows "(\<forall>y \<in> T. \<exists>x \<in> S. f x = y) \<longleftrightarrow> inj_on f S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1015
  (is "?lhs \<longleftrightarrow> ?rhs")
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1016
proof
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1017
  assume h: "?lhs"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1018
  {
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1019
    fix x y
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1020
    assume x: "x \<in> S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1021
    assume y: "y \<in> S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1022
    assume f: "f x = f y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1023
    from x fS have S0: "card S \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1024
      by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1025
    have "x = y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1026
    proof (rule ccontr)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1027
      assume xy: "\<not> ?thesis"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1028
      have th: "card S \<le> card (f ` (S - {y}))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1029
        unfolding c
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1030
      proof (rule card_mono)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1031
        show "finite (f ` (S - {y}))"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1032
          by (simp add: fS)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1033
        show "T \<subseteq> f ` (S - {y})"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1034
          using h xy x y f unfolding subset_eq image_iff
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1035
          by (metis member_remove remove_def)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1036
      qed
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1037
      also have " \<dots> \<le> card (S - {y})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1038
        apply (rule card_image_le)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1039
        using fS by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1040
      also have "\<dots> \<le> card S - 1" using y fS by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1041
      finally show False using S0 by arith
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1042
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1043
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1044
  then show ?rhs
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1045
    unfolding inj_on_def by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1046
next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1047
  assume h: ?rhs
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1048
  have "f ` S = T"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1049
    by (simp add: ST c card_image card_subset_eq fT h)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1050
  then show ?lhs by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1051
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1052
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1053
locale finite_dimensional_vector_space = vector_space +
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1054
  fixes Basis :: "'b set"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1055
  assumes finite_Basis: "finite Basis"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1056
  and independent_Basis: "independent Basis"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1057
  and span_Basis: "span Basis = UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1058
begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1059
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1060
definition "dimension = card Basis"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1061
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1062
lemma finiteI_independent: "independent B \<Longrightarrow> finite B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1063
  using independent_span_bound[OF finite_Basis, of B] by (auto simp: span_Basis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1064
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1065
lemma dim_empty [simp]: "dim {} = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1066
  by (rule dim_unique[OF order_refl]) (auto simp: dependent_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1067
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1068
lemma dim_insert:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1069
  "dim (insert x S) = (if x \<in> span S then dim S else dim S + 1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1070
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1071
  show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1072
  proof (cases "x \<in> span S")
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1073
    case True then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1074
      by (metis dim_span span_redundant)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1075
  next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1076
    case False
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1077
    obtain B where B: "B \<subseteq> span S" "independent B" "span S \<subseteq> span B" "card B = dim (span S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1078
      using basis_exists [of "span S"] by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1079
    have 1: "insert x B \<subseteq> span (insert x S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1080
      by (meson B(1) insertI1 insert_subset order_trans span_base span_mono subset_insertI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1081
    have 2: "span (insert x S) \<subseteq> span (insert x B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1082
      by (metis \<open>B \<subseteq> span S\<close> \<open>span S \<subseteq> span B\<close> span_breakdown_eq span_subspace subsetI subspace_span)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1083
    have 3: "independent (insert x B)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1084
      by (metis B(1-3) independent_insert span_subspace subspace_span False)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1085
    have "dim (span (insert x S)) = Suc (dim S)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1086
      apply (rule dim_unique [OF 1 2 3])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1087
      by (metis B False card_insert_disjoint dim_span finiteI_independent span_base span_eq span_span)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1088
    then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1089
      by (metis False Suc_eq_plus1 dim_span)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1090
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1091
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1092
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1093
lemma dim_singleton [simp]:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1094
  "dim{x} = (if x = 0 then 0 else 1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1095
  by (simp add: dim_insert)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1096
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1097
proposition choose_subspace_of_subspace:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1098
  assumes "n \<le> dim S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1099
  obtains T where "subspace T" "T \<subseteq> span S" "dim T = n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1100
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1101
  have "\<exists>T. subspace T \<and> T \<subseteq> span S \<and> dim T = n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1102
  using assms
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1103
  proof (induction n)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1104
    case 0 then show ?case by (auto intro!: exI[where x="{0}"] span_zero)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1105
  next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1106
    case (Suc n)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1107
    then obtain T where "subspace T" "T \<subseteq> span S" "dim T = n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1108
      by force
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1109
    then show ?case
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1110
    proof (cases "span S \<subseteq> span T")
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1111
      case True
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1112
      have "dim S = dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1113
        apply (rule span_eq_dim [OF subset_antisym [OF True]])
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1114
        by (simp add: \<open>T \<subseteq> span S\<close> span_minimal)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1115
      then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1116
        using Suc.prems \<open>dim T = n\<close> by linarith
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1117
    next
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1118
      case False
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1119
      then obtain y where y: "y \<in> S" "y \<notin> T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1120
        by (meson span_mono subsetI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1121
      then have "span (insert y T) \<subseteq> span S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1122
        by (metis (no_types) \<open>T \<subseteq> span S\<close> subsetD insert_subset span_superset span_mono span_span)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1123
      with \<open>dim T = n\<close>  \<open>subspace T\<close> y show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1124
        apply (rule_tac x="span(insert y T)" in exI)
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1125
        apply (auto simp: dim_insert)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1126
        using span_eq_iff by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1127
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1128
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1129
  with that show ?thesis by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1130
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1131
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1132
lemma basis_subspace_exists:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1133
  assumes "subspace S"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1134
  obtains B where "finite B" "B \<subseteq> S" "independent B" "span B = S" "card B = dim S"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1135
  by (metis assms span_subspace basis_exists finiteI_independent)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1136
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1137
lemma dim_mono: assumes "V \<subseteq> span W" shows "dim V \<le> dim W"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1138
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1139
  obtain B where "independent B" "B \<subseteq> W" "W \<subseteq> span B"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1140
    using maximal_independent_subset[of W] by force
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1141
  with dim_le_card[of V B] assms independent_span_bound[of Basis B] basis_card_eq_dim[of B W]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1142
    span_mono[of B W] span_minimal[OF _ subspace_span, of W B]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1143
  show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1144
    by (auto simp: finite_Basis span_Basis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1145
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1146
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1147
lemma dim_subset: "S \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1148
  using dim_mono[of S T] by (auto intro: span_base)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1149
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1150
lemma dim_eq_0 [simp]:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1151
  "dim S = 0 \<longleftrightarrow> S \<subseteq> {0}"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1152
  by (metis basis_exists card_eq_0_iff dim_span finiteI_independent span_empty subset_empty subset_singletonD)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1153
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1154
lemma dim_UNIV[simp]: "dim UNIV = card Basis"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1155
  using dim_eq_card[of Basis UNIV] by (simp add: independent_Basis span_Basis)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1156
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1157
lemma independent_card_le_dim: assumes "B \<subseteq> V" and "independent B" shows "card B \<le> dim V"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1158
  by (subst dim_eq_card[symmetric, OF refl \<open>independent B\<close>]) (rule dim_subset[OF \<open>B \<subseteq> V\<close>])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1159
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1160
lemma dim_subset_UNIV: "dim S \<le> dimension"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1161
  by (metis dim_subset subset_UNIV dim_UNIV dimension_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1162
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1163
lemma card_ge_dim_independent:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1164
  assumes BV: "B \<subseteq> V"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1165
    and iB: "independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1166
    and dVB: "dim V \<le> card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1167
  shows "V \<subseteq> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1168
proof
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1169
  fix a
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1170
  assume aV: "a \<in> V"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1171
  {
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1172
    assume aB: "a \<notin> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1173
    then have iaB: "independent (insert a B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1174
      using iB aV BV by (simp add: independent_insert)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1175
    from aV BV have th0: "insert a B \<subseteq> V"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1176
      by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1177
    from aB have "a \<notin>B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1178
      by (auto simp add: span_base)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1179
    with independent_card_le_dim[OF th0 iaB] dVB finiteI_independent[OF iB]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1180
    have False by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1181
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1182
  then show "a \<in> span B" by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1183
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1184
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1185
lemma card_le_dim_spanning:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1186
  assumes BV: "B \<subseteq> V"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1187
    and VB: "V \<subseteq> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1188
    and fB: "finite B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1189
    and dVB: "dim V \<ge> card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1190
  shows "independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1191
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1192
  {
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1193
    fix a
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1194
    assume a: "a \<in> B" "a \<in> span (B - {a})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1195
    from a fB have c0: "card B \<noteq> 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1196
      by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1197
    from a fB have cb: "card (B - {a}) = card B - 1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1198
      by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1199
    {
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1200
      fix x
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1201
      assume x: "x \<in> V"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1202
      from a have eq: "insert a (B - {a}) = B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1203
        by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1204
      from x VB have x': "x \<in> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1205
        by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1206
      from span_trans[OF a(2), unfolded eq, OF x']
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1207
      have "x \<in> span (B - {a})" .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1208
    }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1209
    then have th1: "V \<subseteq> span (B - {a})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1210
      by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1211
    have th2: "finite (B - {a})"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1212
      using fB by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1213
    from dim_le_card[OF th1 th2]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1214
    have c: "dim V \<le> card (B - {a})" .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1215
    from c c0 dVB cb have False by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1216
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1217
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1218
    unfolding dependent_def by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1219
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1220
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1221
lemma card_eq_dim: "B \<subseteq> V \<Longrightarrow> card B = dim V \<Longrightarrow> finite B \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1222
  by (metis order_eq_iff card_le_dim_spanning card_ge_dim_independent)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1223
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1224
lemma subspace_dim_equal:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1225
  assumes "subspace S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1226
    and "subspace T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1227
    and "S \<subseteq> T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1228
    and "dim S \<ge> dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1229
  shows "S = T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1230
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1231
  obtain B where B: "B \<le> S" "independent B \<and> S \<subseteq> span B" "card B = dim S"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1232
    using basis_exists[of S] by metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1233
  then have "span B \<subseteq> S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1234
    using span_mono[of B S] span_eq_iff[of S] assms by metis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1235
  then have "span B = S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1236
    using B by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1237
  have "dim S = dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1238
    using assms dim_subset[of S T] by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1239
  then have "T \<subseteq> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1240
    using card_eq_dim[of B T] B finiteI_independent assms by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1241
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1242
    using assms \<open>span B = S\<close> by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1243
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1244
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1245
corollary dim_eq_span:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1246
  shows "\<lbrakk>S \<subseteq> T; dim T \<le> dim S\<rbrakk> \<Longrightarrow> span S = span T"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1247
  by (simp add: span_mono subspace_dim_equal)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1248
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1249
lemma dim_psubset:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1250
  "span S \<subset> span T \<Longrightarrow> dim S < dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1251
  by (metis (no_types, hide_lams) dim_span less_le not_le subspace_dim_equal subspace_span)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1252
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1253
lemma dim_eq_full:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1254
  shows "dim S = dimension \<longleftrightarrow> span S = UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1255
  by (metis dim_eq_span dim_subset_UNIV span_Basis span_span subset_UNIV
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1256
        dim_UNIV dim_span dimension_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1257
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1258
lemma indep_card_eq_dim_span:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1259
  assumes "independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1260
  shows "finite B \<and> card B = dim (span B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1261
  using dim_span_eq_card_independent[OF assms] finiteI_independent[OF assms] by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1262
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1263
text \<open>More general size bound lemmas.\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1264
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1265
lemma independent_bound_general:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1266
  "independent S \<Longrightarrow> finite S \<and> card S \<le> dim S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1267
  by (simp add: dim_eq_card_independent finiteI_independent)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1268
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1269
lemma independent_explicit:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1270
  shows "independent B \<longleftrightarrow> finite B \<and> (\<forall>c. (\<Sum>v\<in>B. c v *s v) = 0 \<longrightarrow> (\<forall>v \<in> B. c v = 0))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1271
  apply (cases "finite B")
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1272
   apply (force simp: dependent_finite)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1273
  using independent_bound_general
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1274
  apply auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1275
  done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1276
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1277
proposition dim_sums_Int:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1278
  assumes "subspace S" "subspace T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1279
  shows "dim {x + y |x y. x \<in> S \<and> y \<in> T} + dim(S \<inter> T) = dim S + dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1280
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1281
  obtain B where B: "B \<subseteq> S \<inter> T" "S \<inter> T \<subseteq> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1282
             and indB: "independent B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1283
             and cardB: "card B = dim (S \<inter> T)"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1284
    using basis_exists by metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1285
  then obtain C D where "B \<subseteq> C" "C \<subseteq> S" "independent C" "S \<subseteq> span C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1286
                    and "B \<subseteq> D" "D \<subseteq> T" "independent D" "T \<subseteq> span D"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1287
    using maximal_independent_subset_extend
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1288
    by (metis Int_subset_iff \<open>B \<subseteq> S \<inter> T\<close> indB)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1289
  then have "finite B" "finite C" "finite D"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1290
    by (simp_all add: finiteI_independent indB independent_bound_general)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1291
  have Beq: "B = C \<inter> D"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1292
    apply (rule sym)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1293
    apply (rule spanning_subset_independent)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1294
    using \<open>B \<subseteq> C\<close> \<open>B \<subseteq> D\<close> apply blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1295
    apply (meson \<open>independent C\<close> independent_mono inf.cobounded1)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1296
    using B \<open>C \<subseteq> S\<close> \<open>D \<subseteq> T\<close> apply auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1297
    done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1298
  then have Deq: "D = B \<union> (D - C)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1299
    by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1300
  have CUD: "C \<union> D \<subseteq> {x + y |x y. x \<in> S \<and> y \<in> T}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1301
    apply safe
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1302
    apply (metis add.right_neutral subsetCE \<open>C \<subseteq> S\<close> \<open>subspace T\<close> set_eq_subset span_zero span_minimal)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1303
    apply (metis add.left_neutral subsetCE \<open>D \<subseteq> T\<close> \<open>subspace S\<close> set_eq_subset span_zero span_minimal)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1304
    done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1305
  have "a v = 0" if 0: "(\<Sum>v\<in>C. a v *s v) + (\<Sum>v\<in>D - C. a v *s v) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1306
                 and v: "v \<in> C \<union> (D-C)" for a v
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1307
  proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1308
    have eq: "(\<Sum>v\<in>D - C. a v *s v) = - (\<Sum>v\<in>C. a v *s v)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1309
      using that add_eq_0_iff by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1310
    have "(\<Sum>v\<in>D - C. a v *s v) \<in> S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1311
      apply (subst eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1312
      apply (rule subspace_neg [OF \<open>subspace S\<close>])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1313
      apply (rule subspace_sum [OF \<open>subspace S\<close>])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1314
      by (meson subsetCE subspace_scale \<open>C \<subseteq> S\<close> \<open>subspace S\<close>)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1315
    moreover have "(\<Sum>v\<in>D - C. a v *s v) \<in> T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1316
      apply (rule subspace_sum [OF \<open>subspace T\<close>])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1317
      by (meson DiffD1 \<open>D \<subseteq> T\<close> \<open>subspace T\<close> subset_eq subspace_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1318
    ultimately have "(\<Sum>v \<in> D-C. a v *s v) \<in> span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1319
      using B by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1320
    then obtain e where e: "(\<Sum>v\<in>B. e v *s v) = (\<Sum>v \<in> D-C. a v *s v)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1321
      using span_finite [OF \<open>finite B\<close>] by force
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1322
    have "\<And>c v. \<lbrakk>(\<Sum>v\<in>C. c v *s v) = 0; v \<in> C\<rbrakk> \<Longrightarrow> c v = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1323
      using \<open>finite C\<close> \<open>independent C\<close> independentD by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1324
    define cc where "cc x = (if x \<in> B then a x + e x else a x)" for x
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1325
    have [simp]: "C \<inter> B = B" "D \<inter> B = B" "C \<inter> - B = C-D" "B \<inter> (D - C) = {}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1326
      using \<open>B \<subseteq> C\<close> \<open>B \<subseteq> D\<close> Beq by blast+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1327
    have f2: "(\<Sum>v\<in>C \<inter> D. e v *s v) = (\<Sum>v\<in>D - C. a v *s v)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1328
      using Beq e by presburger
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1329
    have f3: "(\<Sum>v\<in>C \<union> D. a v *s v) = (\<Sum>v\<in>C - D. a v *s v) + (\<Sum>v\<in>D - C. a v *s v) + (\<Sum>v\<in>C \<inter> D. a v *s v)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1330
      using \<open>finite C\<close> \<open>finite D\<close> sum.union_diff2 by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1331
    have f4: "(\<Sum>v\<in>C \<union> (D - C). a v *s v) = (\<Sum>v\<in>C. a v *s v) + (\<Sum>v\<in>D - C. a v *s v)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1332
      by (meson Diff_disjoint \<open>finite C\<close> \<open>finite D\<close> finite_Diff sum.union_disjoint)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1333
    have "(\<Sum>v\<in>C. cc v *s v) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1334
      using 0 f2 f3 f4
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1335
      apply (simp add: cc_def Beq \<open>finite C\<close> sum.If_cases algebra_simps sum.distrib
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1336
          if_distrib if_distribR)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1337
      apply (simp add: add.commute add.left_commute diff_eq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1338
      done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1339
    then have "\<And>v. v \<in> C \<Longrightarrow> cc v = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1340
      using independent_explicit \<open>independent C\<close> \<open>finite C\<close> by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1341
    then have C0: "\<And>v. v \<in> C - B \<Longrightarrow> a v = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1342
      by (simp add: cc_def Beq) meson
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1343
    then have [simp]: "(\<Sum>x\<in>C - B. a x *s x) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1344
      by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1345
    have "(\<Sum>x\<in>C. a x *s x) = (\<Sum>x\<in>B. a x *s x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1346
    proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1347
      have "C - D = C - B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1348
        using Beq by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1349
      then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1350
        using Beq \<open>(\<Sum>x\<in>C - B. a x *s x) = 0\<close> f3 f4 by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1351
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1352
    with 0 have Dcc0: "(\<Sum>v\<in>D. a v *s v) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1353
      apply (subst Deq)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1354
      by (simp add: \<open>finite B\<close> \<open>finite D\<close> sum_Un)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1355
    then have D0: "\<And>v. v \<in> D \<Longrightarrow> a v = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1356
      using independent_explicit \<open>independent D\<close> \<open>finite D\<close> by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1357
    show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1358
      using v C0 D0 Beq by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1359
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1360
  then have "independent (C \<union> (D - C))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1361
    unfolding independent_explicit
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1362
    using independent_explicit
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1363
    by (simp add: independent_explicit \<open>finite C\<close> \<open>finite D\<close> sum_Un del: Un_Diff_cancel)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1364
  then have indCUD: "independent (C \<union> D)" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1365
  have "dim (S \<inter> T) = card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1366
    by (rule dim_unique [OF B indB refl])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1367
  moreover have "dim S = card C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1368
    by (metis \<open>C \<subseteq> S\<close> \<open>independent C\<close> \<open>S \<subseteq> span C\<close> basis_card_eq_dim)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1369
  moreover have "dim T = card D"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1370
    by (metis \<open>D \<subseteq> T\<close> \<open>independent D\<close> \<open>T \<subseteq> span D\<close> basis_card_eq_dim)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1371
  moreover have "dim {x + y |x y. x \<in> S \<and> y \<in> T} = card(C \<union> D)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1372
    apply (rule dim_unique [OF CUD _ indCUD refl], clarify)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1373
    apply (meson \<open>S \<subseteq> span C\<close> \<open>T \<subseteq> span D\<close> span_add span_superset span_minimal subsetCE subspace_span sup.bounded_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1374
    done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1375
  ultimately show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1376
    using \<open>B = C \<inter> D\<close> [symmetric]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1377
    by (simp add:  \<open>independent C\<close> \<open>independent D\<close> card_Un_Int finiteI_independent)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1378
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1379
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1380
lemma dependent_biggerset_general:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1381
  "(finite S \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1382
  using independent_bound_general[of S] by (metis linorder_not_le)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1383
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1384
lemma subset_le_dim:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1385
  "S \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1386
  by (metis dim_span dim_subset)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1387
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1388
lemma linear_inj_imp_surj:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1389
  assumes lf: "linear scale scale f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1390
    and fi: "inj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1391
  shows "surj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1392
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1393
  interpret lf: linear scale scale f by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1394
  from basis_exists[of UNIV] obtain B
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1395
    where B: "B \<subseteq> UNIV" "independent B" "UNIV \<subseteq> span B" "card B = dim UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1396
    by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1397
  from B(4) have d: "dim UNIV = card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1398
    by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1399
  have th: "UNIV \<subseteq> span (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1400
    apply (rule card_ge_dim_independent)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1401
      apply blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1402
    using B(2) inj_on_subset[OF fi]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1403
     apply (rule lf.independent_inj_on_image)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1404
     apply blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1405
    apply (rule order_eq_refl)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1406
    apply (rule sym)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1407
    unfolding d
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1408
    apply (rule card_image)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1409
    apply (rule subset_inj_on[OF fi])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1410
    apply blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1411
    done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1412
  from th show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1413
    unfolding lf.span_image surj_def
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1414
    using B(3) by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1415
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1416
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1417
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1418
68620
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1419
locale finite_dimensional_vector_space_pair_1 =
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1420
  vs1: finite_dimensional_vector_space s1 B1 + vs2: vector_space s2
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1421
  for s1 :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr "*a" 75)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1422
  and B1 :: "'b set"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1423
  and s2 :: "'a::field \<Rightarrow> 'c::ab_group_add \<Rightarrow> 'c" (infixr "*b" 75)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1424
begin
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1425
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1426
sublocale vector_space_pair s1 s2 by unfold_locales
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1427
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1428
lemma dim_image_eq:
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1429
  assumes lf: "linear s1 s2 f"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1430
    and fi: "inj_on f (vs1.span S)"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1431
  shows "vs2.dim (f ` S) = vs1.dim S"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1432
proof -
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1433
  interpret lf: linear by fact
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1434
  obtain B where B: "B \<subseteq> S" "vs1.independent B" "S \<subseteq> vs1.span B" "card B = vs1.dim S"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1435
    using vs1.basis_exists[of S] by auto
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1436
  then have "vs1.span S = vs1.span B"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1437
    using vs1.span_mono[of B S] vs1.span_mono[of S "vs1.span B"] vs1.span_span[of B] by auto
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1438
  moreover have "card (f ` B) = card B"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1439
    using assms card_image[of f B] subset_inj_on[of f "vs1.span S" B] B vs1.span_superset by auto
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1440
  moreover have "(f ` B) \<subseteq> (f ` S)"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1441
    using B by auto
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1442
  ultimately show ?thesis
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1443
    by (metis B(2) B(4) fi lf.dependent_inj_imageD lf.span_image vs2.dim_eq_card_independent vs2.dim_span)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1444
qed
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1445
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1446
lemma dim_image_le:
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1447
  assumes lf: "linear s1 s2 f"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1448
  shows "vs2.dim (f ` S) \<le> vs1.dim (S)"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1449
proof -
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1450
  from vs1.basis_exists[of S] obtain B where
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1451
    B: "B \<subseteq> S" "vs1.independent B" "S \<subseteq> vs1.span B" "card B = vs1.dim S" by blast
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1452
  from B have fB: "finite B" "card B = vs1.dim S"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1453
    using vs1.independent_bound_general by blast+
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1454
  have "vs2.dim (f ` S) \<le> card (f ` B)"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1455
    apply (rule vs2.span_card_ge_dim)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1456
    using lf B fB
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1457
      apply (auto simp add: module_hom.span_image module_hom.spans_image subset_image_iff
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1458
        linear_iff_module_hom)
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1459
    done
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1460
  also have "\<dots> \<le> vs1.dim S"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1461
    using card_image_le[OF fB(1)] fB by simp
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1462
  finally show ?thesis .
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1463
qed
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1464
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1465
end
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1466
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1467
locale finite_dimensional_vector_space_pair =
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1468
  vs1: finite_dimensional_vector_space s1 B1 + vs2: finite_dimensional_vector_space s2 B2
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1469
  for s1 :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr "*a" 75)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1470
  and B1 :: "'b set"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1471
  and s2 :: "'a::field \<Rightarrow> 'c::ab_group_add \<Rightarrow> 'c" (infixr "*b" 75)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1472
  and B2 :: "'c set"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1473
begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1474
68620
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1475
sublocale finite_dimensional_vector_space_pair_1 ..
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1476
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1477
lemma linear_surjective_imp_injective:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1478
  assumes lf: "linear s1 s2 f" and sf: "surj f" and eq: "vs2.dim UNIV = vs1.dim UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1479
    shows "inj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1480
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1481
  interpret linear s1 s2 f by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1482
  have *: "card (f ` B1) \<le> vs2.dim UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1483
    using vs1.finite_Basis vs1.dim_eq_card[of B1 UNIV] sf
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1484
    by (auto simp: vs1.span_Basis vs1.independent_Basis eq
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1485
        simp del: vs2.dim_UNIV
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1486
        intro!: card_image_le)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1487
  have indep_fB: "vs2.independent (f ` B1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1488
    using vs1.finite_Basis vs1.dim_eq_card[of B1 UNIV] sf *
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1489
    by (intro vs2.card_le_dim_spanning[of "f ` B1" UNIV]) (auto simp: span_image vs1.span_Basis )
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1490
  have "vs2.dim UNIV \<le> card (f ` B1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1491
    unfolding eq sf[symmetric] vs2.dim_span_eq_card_independent[symmetric, OF indep_fB]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1492
      vs2.dim_span
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1493
    by (intro vs2.dim_mono) (auto simp: span_image vs1.span_Basis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1494
  with * have "card (f ` B1) = vs2.dim UNIV" by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1495
  also have "... = card B1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1496
    unfolding eq vs1.dim_UNIV ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1497
  finally have "inj_on f B1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1498
    by (subst inj_on_iff_eq_card[OF vs1.finite_Basis])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1499
  then show "inj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1500
    using inj_on_span_iff_independent_image[OF indep_fB] vs1.span_Basis by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1501
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1502
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1503
lemma linear_injective_imp_surjective:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1504
  assumes lf: "linear s1 s2 f" and sf: "inj f" and eq: "vs2.dim UNIV = vs1.dim UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1505
    shows "surj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1506
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1507
  interpret linear s1 s2 f by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1508
  have *: False if b: "b \<notin> vs2.span (f ` B1)" for b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1509
  proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1510
    have *: "vs2.independent (f ` B1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1511
      using vs1.independent_Basis by (intro independent_injective_image inj_on_subset[OF sf]) auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1512
    have **: "vs2.independent (insert b (f ` B1))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1513
      using b * by (rule vs2.independent_insertI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1514
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1515
    have "b \<notin> f ` B1" using vs2.span_base[of b "f ` B1"] b by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1516
    then have "Suc (card B1) = card (insert b (f ` B1))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1517
      using sf[THEN inj_on_subset, of B1] by (subst card_insert) (auto intro: vs1.finite_Basis simp: card_image)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1518
    also have "\<dots> = vs2.dim (insert b (f ` B1))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1519
      using vs2.dim_eq_card_independent[OF **] by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1520
    also have "vs2.dim (insert b (f ` B1)) \<le> vs2.dim B2"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1521
      by (rule vs2.dim_mono) (auto simp: vs2.span_Basis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1522
    also have "\<dots> = card B1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1523
      using vs1.dim_span[of B1] vs2.dim_span[of B2] unfolding vs1.span_Basis vs2.span_Basis eq 
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1524
        vs1.dim_eq_card_independent[OF vs1.independent_Basis] by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1525
    finally show False by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1526
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1527
  have "f ` UNIV = f ` vs1.span B1" unfolding vs1.span_Basis ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1528
  also have "\<dots> = vs2.span (f ` B1)" unfolding span_image ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1529
  also have "\<dots> = UNIV" using * by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1530
  finally show ?thesis .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1531
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1532
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1533
lemma linear_injective_isomorphism:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1534
  assumes lf: "linear s1 s2 f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1535
    and fi: "inj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1536
    and dims: "vs2.dim UNIV = vs1.dim UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1537
  shows "\<exists>f'. linear s2 s1 f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1538
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1539
  show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1540
    unfolding isomorphism_expand[symmetric]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1541
    using linear_exists_right_inverse_on[OF lf vs1.subspace_UNIV]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1542
      linear_exists_left_inverse_on[OF lf vs1.subspace_UNIV fi]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1543
    apply (auto simp: module_hom_iff_linear)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1544
    subgoal for f' f''
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1545
      apply (rule exI[where x=f''])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1546
      using linear_injective_imp_surjective[OF lf fi dims]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1547
      apply auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1548
      by (metis comp_apply eq_id_iff surj_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1549
    done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1550
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1551
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1552
lemma linear_surjective_isomorphism:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1553
  assumes lf: "linear s1 s2 f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1554
    and sf: "surj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1555
    and dims: "vs2.dim UNIV = vs1.dim UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1556
  shows "\<exists>f'. linear s2 s1 f' \<and> (\<forall>x. f' (f x) = x) \<and> (\<forall>x. f (f' x) = x)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1557
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1558
  show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1559
    unfolding isomorphism_expand[symmetric]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1560
    using linear_exists_right_inverse_on[OF lf vs1.subspace_UNIV]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1561
      linear_exists_left_inverse_on[OF lf vs1.subspace_UNIV]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1562
    using linear_surjective_imp_injective[OF lf sf dims] sf
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1563
    apply (auto simp: module_hom_iff_linear)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1564
    subgoal for f' f''
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1565
      apply (rule exI[where x=f''])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1566
      apply auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1567
      by (metis isomorphism_expand)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1568
    done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1569
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1570
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1571
lemma basis_to_basis_subspace_isomorphism:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1572
  assumes s: "vs1.subspace S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1573
    and t: "vs2.subspace T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1574
    and d: "vs1.dim S = vs2.dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1575
    and B: "B \<subseteq> S" "vs1.independent B" "S \<subseteq> vs1.span B" "card B = vs1.dim S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1576
    and C: "C \<subseteq> T" "vs2.independent C" "T \<subseteq> vs2.span C" "card C = vs2.dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1577
  shows "\<exists>f. linear s1 s2 f \<and> f ` B = C \<and> f ` S = T \<and> inj_on f S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1578
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1579
  from B have fB: "finite B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1580
    by (simp add: vs1.finiteI_independent)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1581
  from C have fC: "finite C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1582
    by (simp add: vs2.finiteI_independent)
68620
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68412
diff changeset
  1583
  from finite_basis_to_basis_subspace_isomorphism[OF s t d fB B fC C] show ?thesis .
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1584
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1585
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1586
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1587
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1588
context finite_dimensional_vector_space begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1589
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1590
lemma linear_surj_imp_inj:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1591
  assumes lf: "linear scale scale f" and sf: "surj f"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1592
  shows "inj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1593
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1594
  interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1595
  let ?U = "UNIV :: 'b set"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1596
  from basis_exists[of ?U] obtain B
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1597
    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" and d: "card B = dim ?U"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1598
    by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1599
  {
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1600
    fix x
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1601
    assume x: "x \<in> span B" and fx: "f x = 0"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1602
    from B(2) have fB: "finite B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1603
      using finiteI_independent by auto
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1604
    have Uspan: "UNIV \<subseteq> span (f ` B)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1605
      by (simp add: B(3) lf linear_spanning_surjective_image sf)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1606
    have fBi: "independent (f ` B)"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1607
    proof (rule card_le_dim_spanning)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1608
      show "card (f ` B) \<le> dim ?U"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1609
        using card_image_le d fB by fastforce
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1610
    qed (use fB Uspan in auto)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1611
    have th0: "dim ?U \<le> card (f ` B)"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1612
      by (rule span_card_ge_dim) (use Uspan fB in auto)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1613
    moreover have "card (f ` B) \<le> card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1614
      by (rule card_image_le, rule fB)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1615
    ultimately have th1: "card B = card (f ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1616
      unfolding d by arith
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1617
    have fiB: "inj_on f B"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1618
      by (simp add: eq_card_imp_inj_on fB th1)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1619
    from linear_indep_image_lemma[OF lf fB fBi fiB x] fx
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1620
    have "x = 0" by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1621
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1622
  then show ?thesis
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1623
    unfolding linear_inj_iff_eq_0[OF lf] using B(3) by blast
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1624
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1625
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1626
lemma linear_inverse_left:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1627
  assumes lf: "linear scale scale f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1628
    and lf': "linear scale scale f'"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1629
  shows "f \<circ> f' = id \<longleftrightarrow> f' \<circ> f = id"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1630
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1631
  {
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1632
    fix f f':: "'b \<Rightarrow> 'b"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1633
    assume lf: "linear scale scale f" "linear scale scale f'"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1634
    assume f: "f \<circ> f' = id"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1635
    from f have sf: "surj f"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1636
      by (auto simp add: o_def id_def surj_def) metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1637
    interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1638
    from linear_surjective_isomorphism[OF lf(1) sf] lf f
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1639
    have "f' \<circ> f = id"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1640
      unfolding fun_eq_iff o_def id_def by metis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1641
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1642
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1643
    using lf lf' by metis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1644
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1645
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1646
lemma left_inverse_linear:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1647
  assumes lf: "linear scale scale f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1648
    and gf: "g \<circ> f = id"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1649
  shows "linear scale scale g"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1650
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1651
  from gf have fi: "inj f"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1652
    by (auto simp add: inj_on_def o_def id_def fun_eq_iff) metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1653
  interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1654
  from linear_injective_isomorphism[OF lf fi]
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1655
  obtain h :: "'b \<Rightarrow> 'b" where "linear scale scale h" and h: "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1656
    by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1657
  have "h = g"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
  1658
    by (metis gf h isomorphism_expand left_right_inverse_eq)
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1659
  with \<open>linear scale scale h\<close> show ?thesis by blast
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1660
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1661
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1662
lemma inj_linear_imp_inv_linear:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1663
  assumes "linear scale scale f" "inj f" shows "linear scale scale (inv f)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1664
  using assms inj_iff left_inverse_linear by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1665
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1666
lemma right_inverse_linear:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1667
  assumes lf: "linear scale scale f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1668
    and gf: "f \<circ> g = id"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1669
  shows "linear scale scale g"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1670
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1671
  from gf have fi: "surj f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1672
    by (auto simp add: surj_def o_def id_def) metis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1673
  interpret finite_dimensional_vector_space_pair scale Basis scale Basis by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1674
  from linear_surjective_isomorphism[OF lf fi]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1675
  obtain h:: "'b \<Rightarrow> 'b" where h: "linear scale scale h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1676
    by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1677
  have "h = g"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1678
    apply (rule ext)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1679
    using gf h(2,3)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1680
    apply (simp add: o_def id_def fun_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1681
    apply metis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1682
    done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1683
  with h(1) show ?thesis by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1684
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1685
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1686
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1687
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1688
context finite_dimensional_vector_space_pair begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1689
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1690
lemma subspace_isomorphism:
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1691
  assumes s: "vs1.subspace S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1692
    and t: "vs2.subspace T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1693
    and d: "vs1.dim S = vs2.dim T"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1694
  shows "\<exists>f. linear s1 s2 f \<and> f ` S = T \<and> inj_on f S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1695
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1696
  from vs1.basis_exists[of S] vs1.finiteI_independent
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1697
  obtain B where B: "B \<subseteq> S" "vs1.independent B" "S \<subseteq> vs1.span B" "card B = vs1.dim S" and fB: "finite B"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1698
    by metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1699
  from vs2.basis_exists[of T] vs2.finiteI_independent
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1700
  obtain C where C: "C \<subseteq> T" "vs2.independent C" "T \<subseteq> vs2.span C" "card C = vs2.dim T" and fC: "finite C"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
  1701
    by metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1702
  from B(4) C(4) card_le_inj[of B C] d
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1703
  obtain f where f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1704
    by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1705
  from linear_independent_extend[OF B(2)]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1706
  obtain g where g: "linear s1 s2 g" "\<forall>x\<in> B. g x = f x"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1707
    by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1708
  interpret g: linear s1 s2 g by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1709
  from inj_on_iff_eq_card[OF fB, of f] f(2) have "card (f ` B) = card B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1710
    by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1711
  with B(4) C(4) have ceq: "card (f ` B) = card C"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1712
    using d by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1713
  have "g ` B = f ` B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1714
    using g(2) by (auto simp add: image_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1715
  also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1716
  finally have gBC: "g ` B = C" .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1717
  have gi: "inj_on g B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1718
    using f(2) g(2) by (auto simp add: inj_on_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1719
  note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1720
  {
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1721
    fix x y
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1722
    assume x: "x \<in> S" and y: "y \<in> S" and gxy: "g x = g y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1723
    from B(3) x y have x': "x \<in> vs1.span B" and y': "y \<in> vs1.span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1724
      by blast+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1725
    from gxy have th0: "g (x - y) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1726
      by (simp add: linear_diff g)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1727
    have th1: "x - y \<in> vs1.span B"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1728
      using x' y' by (metis vs1.span_diff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1729
    have "x = y"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1730
      using g0[OF th1 th0] by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1731
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1732
  then have giS: "inj_on g S"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1733
    unfolding inj_on_def by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1734
  from vs1.span_subspace[OF B(1,3) s] have "g ` S = vs2.span (g ` B)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1735
    by (simp add: module_hom.span_image[OF g(1)[unfolded linear_iff_module_hom]])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1736
  also have "\<dots> = vs2.span C" unfolding gBC ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1737
  also have "\<dots> = T" using vs2.span_subspace[OF C(1,3) t] .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1738
  finally have gS: "g ` S = T" .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1739
  from g(1) gS giS show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1740
    by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1741
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1742
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1743
end
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1744
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1745
hide_const (open) linear
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1746
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1747
end