src/HOL/Algebra/Embedded_Algebras.thy
author paulson <lp15@cam.ac.uk>
Mon, 02 Jul 2018 16:20:03 +0100
changeset 68571 2a6e258bfd66
parent 68569 c64319959bab
child 68582 b9b9e2985878
permissions -rw-r--r--
fixed (?) LaTeX presentation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
(* ************************************************************************** *)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
(* Title:      Embedded_Algebras.thy                                          *)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
(* Author:     Paulo Emílio de Vilhena                                        *)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
(* ************************************************************************** *)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
theory Embedded_Algebras
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
  imports Subrings Generated_Groups
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
begin
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
section \<open>Definitions\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
locale embedded_algebra =
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
  K?: subfield K R + R?: ring R for K :: "'a set" and R :: "('a, 'b) ring_scheme" (structure)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
definition (in ring) line_extension :: "'a set \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> 'a set"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  where "line_extension K a E = (K #> a) <+>\<^bsub>R\<^esub> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
fun (in ring) Span :: "'a set \<Rightarrow> 'a list \<Rightarrow> 'a set"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
  where "Span K Us = foldr (line_extension K) Us { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
fun (in ring) combine :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
  where
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
    "combine (k # Ks) (u # Us) = (k \<otimes> u) \<oplus> (combine Ks Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  | "combine Ks Us = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
inductive (in ring) independent :: "'a set \<Rightarrow> 'a list \<Rightarrow> bool"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
  where
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
    li_Nil [simp, intro]: "independent K []"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  | li_Cons: "\<lbrakk> u \<in> carrier R; u \<notin> Span K Us; independent K Us \<rbrakk> \<Longrightarrow> independent K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
inductive (in ring) dimension :: "nat \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
  where
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
    zero_dim [simp, intro]: "dimension 0 K { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
   | Suc_dim: "\<lbrakk> v \<in> carrier R; v \<notin> E; dimension n K E \<rbrakk> \<Longrightarrow> dimension (Suc n) K (line_extension K v E)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
subsubsection \<open>Syntactic Definitions\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
abbreviation (in ring) dependent ::  "'a set \<Rightarrow> 'a list \<Rightarrow> bool"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
  where "dependent K U \<equiv> \<not> independent K U"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
definition over :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" (infixl "over" 65)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
  where "f over a = f a"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
context ring
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
begin
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
subsection \<open>Basic Properties - First Part\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
lemma combine_in_carrier [simp, intro]:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
  "\<lbrakk> set Ks \<subseteq> carrier R; set Us \<subseteq> carrier R \<rbrakk> \<Longrightarrow> combine Ks Us \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
  by (induct Ks Us rule: combine.induct) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
lemma combine_r_distr:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
  "\<lbrakk> set Ks \<subseteq> carrier R; set Us \<subseteq> carrier R \<rbrakk> \<Longrightarrow>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
     k \<in> carrier R \<Longrightarrow> k \<otimes> (combine Ks Us) = combine (map ((\<otimes>) k) Ks) Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
  by (induct Ks Us rule: combine.induct) (auto simp add: m_assoc r_distr)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
lemma combine_l_distr:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
  "\<lbrakk> set Ks \<subseteq> carrier R; set Us \<subseteq> carrier R \<rbrakk> \<Longrightarrow>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
     u \<in> carrier R \<Longrightarrow> (combine Ks Us) \<otimes> u = combine Ks (map (\<lambda>u'. u' \<otimes> u) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  by (induct Ks Us rule: combine.induct) (auto simp add: m_assoc l_distr)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
lemma combine_eq_foldr:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
  "combine Ks Us = foldr (\<lambda>(k, u). \<lambda>l. (k \<otimes> u) \<oplus> l) (zip Ks Us) \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
  by (induct Ks Us rule: combine.induct) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
lemma combine_replicate:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
  "set Us \<subseteq> carrier R \<Longrightarrow> combine (replicate (length Us) \<zero>) Us = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
  by (induct Us) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
lemma combine_append:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
  assumes "length Ks = length Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
    and "set Ks  \<subseteq> carrier R" "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
    and "set Ks' \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  shows "(combine Ks Us) \<oplus> (combine Ks' Vs) = combine (Ks @ Ks') (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
proof (induct Ks arbitrary: Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
  case Nil thus ?case by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  case (Cons k Ks)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
  then obtain u Us' where Us: "Us = u # Us'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
    by (metis length_Suc_conv)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
  hence u: "u \<in> carrier R" and Us': "set Us' \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
    using Cons(4) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
  then show ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
    using combine_in_carrier[OF _ Us', of Ks] Cons
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
          combine_in_carrier[OF Cons(5-6)] unfolding Us
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
    by (auto, simp add: add.m_assoc)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
lemma combine_add:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
  assumes "length Ks = length Us" and "length Ks' = length Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
    and "set Ks  \<subseteq> carrier R" "set Ks'  \<subseteq> carrier R" "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
  shows "(combine Ks Us) \<oplus> (combine Ks' Us) = combine (map2 (\<oplus>) Ks Ks') Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
proof (induct Us arbitrary: Ks Ks')
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
  case Nil thus ?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  case (Cons u Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
  then obtain c c' Cs Cs' where Ks: "Ks = c # Cs" and Ks': "Ks' = c' # Cs'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
    by (metis length_Suc_conv)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
  hence in_carrier:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
    "c  \<in> carrier R" "set Cs  \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
    "c' \<in> carrier R" "set Cs' \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
    "u  \<in> carrier R" "set Us  \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
    using Cons(4-6) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
  hence lc_in_carrier: "combine Cs Us \<in> carrier R" "combine Cs' Us \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
    using combine_in_carrier by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  have "combine Ks (u # Us) \<oplus> combine Ks' (u # Us) =
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
        ((c \<otimes> u) \<oplus> combine Cs Us) \<oplus> ((c' \<otimes> u) \<oplus> combine Cs' Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    unfolding Ks Ks' by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
  also have " ... = ((c \<oplus> c') \<otimes> u \<oplus> (combine Cs Us \<oplus> combine Cs' Us))"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
    using lc_in_carrier in_carrier(1,3,5) by (simp add: l_distr ring_simprules(7,22))
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
  also have " ... = combine (map2 (\<oplus>) Ks Ks') (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
    using Cons unfolding Ks Ks' by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
  finally show ?case .
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
68571
2a6e258bfd66 fixed (?) LaTeX presentation
paulson <lp15@cam.ac.uk>
parents: 68569
diff changeset
   125
subsection \<open>Some Basic Properties of Linear Independence\<close>
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
lemma independent_in_carrier: "independent K Us \<Longrightarrow> set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  by (induct Us rule: independent.induct) (simp_all)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
lemma independent_backwards:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
  "independent K (u # Us) \<Longrightarrow> u \<notin> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
  "independent K (u # Us) \<Longrightarrow> independent K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
  "independent K (u # Us) \<Longrightarrow> u \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
  by (cases rule: independent.cases, auto)+
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
text \<open>Now, we fix K, a subfield of the ring. Many lemmas would also be true for weaker
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
      structures, but our interest is to work with subfields, so generalization could
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
      be the subjuct of a future work.\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
context
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
  fixes K :: "'a set" assumes K: "subfield K R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
begin
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
subsection \<open>Basic Properties - Second Part\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
lemmas subring_props [simp] =
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
  subringE[OF subfieldE(1)[OF K]]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
lemma line_extension_mem_iff: "u \<in> line_extension K a E \<longleftrightarrow> (\<exists>k \<in> K. \<exists>v \<in> E. u = k \<otimes> a \<oplus> v)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  unfolding line_extension_def set_add_def'[of R "K #> a" E] unfolding r_coset_def by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
lemma line_extension_is_subgroup:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
  assumes "subgroup E (add_monoid R)" "a \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
  shows "subgroup (line_extension K a E) (add_monoid R)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
proof (rule add.subgroupI)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
  show "line_extension K a E \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
    by (simp add: assms add.subgroupE(1) line_extension_def r_coset_subset_G set_add_closed)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
  have "\<zero> = \<zero> \<otimes> a \<oplus> \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
    using assms(2) by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
  hence "\<zero> \<in> line_extension K a E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
    using line_extension_mem_iff subgroup.one_closed[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
  thus "line_extension K a E \<noteq> {}" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
  fix u1 u2
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
  assume "u1 \<in> line_extension K a E" and "u2 \<in> line_extension K a E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  then obtain k1 k2 v1 v2
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
    where u1: "k1 \<in> K" "v1 \<in> E" "u1 = (k1 \<otimes> a) \<oplus> v1"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
      and u2: "k2 \<in> K" "v2 \<in> E" "u2 = (k2 \<otimes> a) \<oplus> v2"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
      and in_carr: "k1 \<in> carrier R" "v1 \<in> carrier R" "k2 \<in> carrier R" "v2 \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
    using line_extension_mem_iff by (meson add.subgroupE(1)[OF assms(1)] subring_props(1) subsetCE)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
  hence "u1 \<oplus> u2 = ((k1 \<oplus> k2) \<otimes> a) \<oplus> (v1 \<oplus> v2)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
    using assms(2) by algebra
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
  moreover have "k1 \<oplus> k2 \<in> K" and "v1 \<oplus> v2 \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
    using add.subgroupE(4)[OF assms(1)] u1 u2 by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  ultimately show "u1 \<oplus> u2 \<in> line_extension K a E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
    using line_extension_mem_iff by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
  have "\<ominus> u1 = ((\<ominus> k1) \<otimes> a) \<oplus> (\<ominus> v1)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
    using in_carr(1-2) u1(3) assms(2) by algebra
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
  moreover have "\<ominus> k1 \<in> K" and "\<ominus> v1 \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
    using add.subgroupE(3)[OF assms(1)] u1 by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  ultimately show "(\<ominus> u1) \<in> line_extension K a E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
    using line_extension_mem_iff by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
corollary Span_is_add_subgroup:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
  "set Us \<subseteq> carrier R \<Longrightarrow> subgroup (Span K Us) (add_monoid R)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
  using line_extension_is_subgroup add.normal_invE(1)[OF add.one_is_normal] by (induct Us) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
lemma line_extension_smult_closed:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
  assumes "\<And>k v. \<lbrakk> k \<in> K; v \<in> E \<rbrakk> \<Longrightarrow> k \<otimes> v \<in> E" and "E \<subseteq> carrier R" "a \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
  shows "\<And>k u. \<lbrakk> k \<in> K; u \<in> line_extension K a E \<rbrakk> \<Longrightarrow> k \<otimes> u \<in> line_extension K a E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
  fix k u assume A: "k \<in> K" "u \<in> line_extension K a E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
  then obtain k' v'
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
    where u: "k' \<in> K" "v' \<in> E" "u = k' \<otimes> a \<oplus> v'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
      and in_carr: "k \<in> carrier R" "k' \<in> carrier R" "v' \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
    using line_extension_mem_iff assms(2) by (meson subring_props(1) subsetCE)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
  hence "k \<otimes> u = (k \<otimes> k') \<otimes> a \<oplus> (k \<otimes> v')"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
    using assms(3) by algebra
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
  thus "k \<otimes> u \<in> line_extension K a E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
    using assms(1)[OF A(1) u(2)] line_extension_mem_iff u(1) A(1) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
lemma Span_subgroup_props [simp]:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
  assumes "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
  shows "Span K Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
    and "\<zero> \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
    and "\<And>v1 v2. \<lbrakk> v1 \<in> Span K Us; v2 \<in> Span K Us \<rbrakk> \<Longrightarrow> (v1 \<oplus> v2) \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
    and "\<And>v. v \<in> Span K Us \<Longrightarrow> (\<ominus> v) \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
  using add.subgroupE subgroup.one_closed[of _ "add_monoid R"]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
        Span_is_add_subgroup[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
lemma Span_smult_closed [simp]:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
  assumes "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
  shows "\<And>k v. \<lbrakk> k \<in> K; v \<in> Span K Us \<rbrakk> \<Longrightarrow> k \<otimes> v \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
proof (induct Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
  case Nil thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
    using r_null subring_props(1) by (auto, blast)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
  case Cons thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
    using Span_subgroup_props(1) line_extension_smult_closed by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
lemma Span_m_inv_simprule [simp]:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
  assumes "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
  shows "\<lbrakk> k \<in> K - { \<zero> }; a \<in> carrier R \<rbrakk> \<Longrightarrow> k \<otimes> a \<in> Span K Us \<Longrightarrow> a \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
  assume k: "k \<in> K - { \<zero> }" and a: "a \<in> carrier R" and ka: "k \<otimes> a \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
  have inv_k: "inv k \<in> K" "inv k \<otimes> k = \<one>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
    using subfield_m_inv[OF K k] by simp+
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
  hence "inv k \<otimes> (k \<otimes> a) \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
    using Span_smult_closed[OF assms _ ka] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
  thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
    using inv_k subring_props(1)a k by (smt DiffD1 l_one m_assoc set_rev_mp)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
subsection \<open>Span as Linear Combinations\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
text \<open>We show that Span is the set of linear combinations\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
lemma line_extension_of_combine_set:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
  assumes "u \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
  shows "line_extension K u { combine Ks Us | Ks. set Ks \<subseteq> K } = 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
                { combine Ks (u # Us) | Ks. set Ks \<subseteq> K }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
  (is "?line_extension = ?combinations")
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
  show "?line_extension \<subseteq> ?combinations"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
    fix v assume "v \<in> ?line_extension"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
    then obtain k Ks
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
      where "k \<in> K" "set Ks \<subseteq> K" and "v = combine (k # Ks) (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
      using line_extension_mem_iff by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
    thus "v \<in> ?combinations"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
      by (metis (mono_tags, lifting) insert_subset list.simps(15) mem_Collect_eq)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
  show "?combinations \<subseteq> ?line_extension"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
    fix v assume "v \<in> ?combinations"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
    then obtain Ks where v: "set Ks \<subseteq> K" "v = combine Ks (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
      by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
    thus "v \<in> ?line_extension"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
    proof (cases Ks)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
      case Cons thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
        using v line_extension_mem_iff by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
    next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
      case Nil
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
      hence "v = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
        using v by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
      moreover have "combine [] Us = \<zero>" by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
      hence "\<zero> \<in> { combine Ks Us | Ks. set Ks \<subseteq> K }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
        by (metis (mono_tags, lifting) local.Nil mem_Collect_eq v(1))
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
      hence "(\<zero> \<otimes> u) \<oplus> \<zero> \<in> ?line_extension"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
        using line_extension_mem_iff subring_props(2) by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
      hence "\<zero> \<in> ?line_extension"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
        using assms by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
      ultimately show ?thesis by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
lemma Span_eq_combine_set:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
  assumes "set Us \<subseteq> carrier R" shows "Span K Us = { combine Ks Us | Ks. set Ks \<subseteq> K }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
  using assms line_extension_of_combine_set
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
  by (induct Us) (auto, metis empty_set empty_subsetI)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
lemma line_extension_of_combine_set_length_version:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
  assumes "u \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
  shows "line_extension K u { combine Ks Us | Ks. length Ks = length Us \<and> set Ks \<subseteq> K } = 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
                      { combine Ks (u # Us) | Ks. length Ks = length (u # Us) \<and> set Ks \<subseteq> K }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
  (is "?line_extension = ?combinations")
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
  show "?line_extension \<subseteq> ?combinations"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
    fix v assume "v \<in> ?line_extension"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
    then obtain k Ks
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
      where "v = combine (k # Ks) (u # Us)" "length (k # Ks) = length (u # Us)" "set (k # Ks) \<subseteq> K"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
      using line_extension_mem_iff by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
    thus "v \<in> ?combinations" by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
  show "?combinations \<subseteq> ?line_extension"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
    fix c assume "c \<in> ?combinations"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
    then obtain Ks where c: "c = combine Ks (u # Us)" "length Ks = length (u # Us)" "set Ks \<subseteq> K"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
      by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
    then obtain k Ks' where k: "Ks = k # Ks'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
      by (metis length_Suc_conv)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
    thus "c \<in> ?line_extension"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
      using c line_extension_mem_iff unfolding k by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
lemma Span_eq_combine_set_length_version:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
  assumes "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
  shows "Span K Us = { combine Ks Us | Ks. length Ks = length Us \<and> set Ks \<subseteq> K }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
  using assms line_extension_of_combine_set_length_version by (induct Us) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
subsubsection \<open>Corollaries\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
corollary Span_mem_iff:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
  assumes "set Us \<subseteq> carrier R" and "a \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
  shows "a \<in> Span K Us \<longleftrightarrow> (\<exists>k \<in> K - { \<zero> }. \<exists>Ks. set Ks \<subseteq> K \<and> combine (k # Ks) (a # Us) = \<zero>)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
         (is "?in_Span \<longleftrightarrow> ?exists_combine")
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
proof 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
  assume "?in_Span"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
  then obtain Ks where Ks: "set Ks \<subseteq> K" "a = combine Ks Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
    using Span_eq_combine_set[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
  hence "((\<ominus> \<one>) \<otimes> a) \<oplus> a = combine ((\<ominus> \<one>) # Ks) (a # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
    by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
  moreover have "((\<ominus> \<one>) \<otimes> a) \<oplus> a = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
    using assms(2) l_minus l_neg by auto  
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
  moreover have "\<ominus> \<one> \<noteq> \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
    using subfieldE(6)[OF K] l_neg by force 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
  ultimately show "?exists_combine"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
    using subring_props(3,5) Ks(1) by (force simp del: combine.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
  assume "?exists_combine"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
  then obtain k Ks
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
    where k: "k \<in> K" "k \<noteq> \<zero>" and Ks: "set Ks \<subseteq> K" and a: "(k \<otimes> a) \<oplus> combine Ks Us = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
    by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
  hence "combine Ks Us \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
    using Span_eq_combine_set[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
  hence "k \<otimes> a \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
    using Span_subgroup_props[OF assms(1)] k Ks a
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
    by (metis (no_types, lifting) assms(2) contra_subsetD m_closed minus_equality subring_props(1))
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
  thus "?in_Span"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
    using Span_m_inv_simprule[OF assms(1) _ assms(2), of k] k by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
corollary Span_mem_iff_length_version:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
  assumes "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
  shows "a \<in> Span K Us \<longleftrightarrow> (\<exists>Ks. set Ks \<subseteq> K \<and> length Ks = length Us \<and> a = combine Ks Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
  using Span_eq_combine_set_length_version[OF assms] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
68571
2a6e258bfd66 fixed (?) LaTeX presentation
paulson <lp15@cam.ac.uk>
parents: 68569
diff changeset
   365
subsection \<open>Span as the minimal subgroup that contains @{term"K <#> (set Us)"}\<close>
68569
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
text \<open>Now we show the link between Span and Group.generate\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
lemma mono_Span:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
  assumes "set Us \<subseteq> carrier R" and "u \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
  shows "Span K Us \<subseteq> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
  fix v assume v: "v \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
  hence "(\<zero> \<otimes> u) \<oplus> v \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
    using line_extension_mem_iff by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
  thus "v \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
    using Span_subgroup_props(1)[OF assms(1)] assms(2) v
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
    by (auto simp del: Span.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
lemma Span_min:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
  assumes "set Us \<subseteq> carrier R" and "subgroup E (add_monoid R)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
  shows "K <#> (set Us) \<subseteq> E \<Longrightarrow> Span K Us \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
  assume "K <#> (set Us) \<subseteq> E" show "Span K Us \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
    fix v assume "v \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
    then obtain Ks where v: "set Ks \<subseteq> K" "v = combine Ks Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
      using Span_eq_combine_set[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
    from \<open>set Ks \<subseteq> K\<close> \<open>set Us \<subseteq> carrier R\<close> and \<open>K <#> (set Us) \<subseteq> E\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
    show "v \<in> E" unfolding v(2)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
    proof (induct Ks Us rule: combine.induct)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
      case (1 k Ks u Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
      hence "k \<in> K" and "u \<in> set (u # Us)" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
      hence "k \<otimes> u \<in> E" 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
        using 1(4) unfolding set_mult_def by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
      moreover have "K <#> set Us \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
        using 1(4) unfolding set_mult_def by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
      hence "combine Ks Us \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
        using 1 by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
      ultimately show ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
        using add.subgroupE(4)[OF assms(2)] by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
    next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
      case "2_1" thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
        using subgroup.one_closed[OF assms(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
    next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
      case  "2_2" thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
        using subgroup.one_closed[OF assms(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
lemma Span_eq_generate:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
  assumes "set Us \<subseteq> carrier R" shows "Span K Us = generate (add_monoid R) (K <#> (set Us))"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
proof (rule add.generateI)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
  show "K <#> set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
    using subring_props(1) assms unfolding set_mult_def by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
  show "subgroup (Span K Us) (add_monoid R)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
    using Span_is_add_subgroup[OF assms] .
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
  show "\<And>E. \<lbrakk> subgroup E (add_monoid R); K <#> set Us \<subseteq> E \<rbrakk> \<Longrightarrow> Span K Us \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
    using Span_min assms by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
  show "K <#> set Us \<subseteq> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
  proof (induct Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
    case Nil thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
      unfolding set_mult_def by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
  next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
    case (Cons u Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
    have "K <#> set (u # Us) = (K <#> { u }) \<union> (K <#> set Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
      unfolding set_mult_def by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
    moreover have "\<And>k. k \<in> K \<Longrightarrow> k \<otimes> u \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
    proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
      fix k assume k: "k \<in> K"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
      hence "combine [ k ] (u # Us) \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
        using Span_eq_combine_set[OF Cons(2)] by (auto simp del: combine.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
      moreover have "k \<in> carrier R" and "u \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
        using Cons(2) k subring_props(1) by (blast, auto) 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
      ultimately show "k \<otimes> u \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
        by (auto simp del: Span.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
    hence "K <#> { u } \<subseteq> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
      unfolding set_mult_def by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
    moreover have "K <#> set Us \<subseteq> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
      using mono_Span[of Us u] Cons by (auto simp del: Span.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
    ultimately show ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
      using Cons by (auto simp del: Span.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
subsubsection \<open>Corollaries\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
corollary Span_same_set:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
  assumes "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
  shows "set Us = set Vs \<Longrightarrow> Span K Us = Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
  using Span_eq_generate assms by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
corollary Span_incl: "set Us \<subseteq> carrier R \<Longrightarrow> K <#> (set Us) \<subseteq> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
  using Span_eq_generate generate.incl[of _ _ "add_monoid R"] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
corollary Span_base_incl: "set Us \<subseteq> carrier R \<Longrightarrow> set Us \<subseteq> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
  assume A: "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
  hence "{ \<one> } <#> set Us = set Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
    unfolding set_mult_def by force
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
  moreover have "{ \<one> } <#> set Us \<subseteq> K <#> set Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
    using subring_props(3) unfolding set_mult_def by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
  ultimately show ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
    using Span_incl[OF A] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
corollary mono_Span_sublist:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
  assumes "set Us \<subseteq> set Vs" "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
  shows "Span K Us \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
  using add.mono_generate[OF mono_set_mult[OF _ assms(1), of K K R]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
        set_mult_closed[OF subring_props(1) assms(2)]]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
        Span_eq_generate[OF assms(2)] Span_eq_generate[of Us] assms by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
corollary mono_Span_append:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
  assumes "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
  shows "Span K Us \<subseteq> Span K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
    and "Span K Us \<subseteq> Span K (Vs @ Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
  using mono_Span_sublist[of Us "Us @ Vs"] assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
        Span_same_set[of "Us @ Vs" "Vs @ Us"] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
corollary mono_Span_subset:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
  assumes "set Us \<subseteq> Span K Vs" "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
  shows "Span K Us \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
proof (rule Span_min[OF _ Span_is_add_subgroup[OF assms(2)]])
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
  show "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
    using Span_subgroup_props(1)[OF assms(2)] assms by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
  show "K <#> set Us \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
    using Span_smult_closed[OF assms(2)] assms(1) unfolding set_mult_def by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   497
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
lemma Span_strict_incl:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
  assumes "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
  shows "Span K Us \<subset> Span K Vs \<Longrightarrow> (\<exists>v \<in> set Vs. v \<notin> Span K Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
  assume "Span K Us \<subset> Span K Vs" show "\<exists>v \<in> set Vs. v \<notin> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
  proof (rule ccontr)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
    assume "\<not> (\<exists>v \<in> set Vs. v \<notin> Span K Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
    hence "Span K Vs \<subseteq> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
      using mono_Span_subset[OF _ assms(1), of Vs] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
    from \<open>Span K Us \<subset> Span K Vs\<close> and \<open>Span K Vs \<subseteq> Span K Us\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
    show False by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
lemma Span_append_eq_set_add:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
  assumes "set Us \<subseteq> carrier R" and "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
  shows "Span K (Us @ Vs) = (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
proof (induct Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
  case Nil thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
    using Span_subgroup_props(1)[OF Nil(2)] unfolding set_add_def' by force
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
  case (Cons u Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
  hence in_carrier:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
    "u \<in> carrier R" "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
    by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
  have "line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs) = (Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
    show "line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs) \<subseteq> (Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
    proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
      fix v assume "v \<in> line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
      then obtain k u' v'
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
        where v: "k \<in> K" "u' \<in> Span K Us" "v' \<in> Span K Vs" "v = k \<otimes> u \<oplus> (u' \<oplus> v')"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
        using line_extension_mem_iff[of v u "Span K Us <+>\<^bsub>R\<^esub> Span K Vs"]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
        unfolding set_add_def' by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
      hence "v = (k \<otimes> u \<oplus> u') \<oplus> v'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
        using in_carrier(2-3)[THEN Span_subgroup_props(1)] in_carrier(1) subring_props(1)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
        by (metis (no_types, lifting) rev_subsetD ring_simprules(7) semiring_simprules(3))
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
      moreover have "k \<otimes> u \<oplus> u' \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
        using line_extension_mem_iff v(1-2) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
      ultimately show "v \<in> Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
        unfolding set_add_def' using v(3) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
  next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
    show "Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs \<subseteq> line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
    proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
      fix v assume "v \<in> Span K (u # Us) <+>\<^bsub>R\<^esub> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
      then obtain k u' v'
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
        where v: "k \<in> K" "u' \<in> Span K Us" "v' \<in> Span K Vs" "v = (k \<otimes> u \<oplus> u') \<oplus> v'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
        using line_extension_mem_iff[of _ u "Span K Us"] unfolding set_add_def' by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
      hence "v = (k \<otimes> u) \<oplus> (u' \<oplus> v')"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
        using in_carrier(2-3)[THEN Span_subgroup_props(1)] in_carrier(1) subring_props(1)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
        by (metis (no_types, lifting) rev_subsetD ring_simprules(5,7))
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
      thus "v \<in> line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   554
        using line_extension_mem_iff[of "(k \<otimes> u) \<oplus> (u' \<oplus> v')" u "Span K Us <+>\<^bsub>R\<^esub> Span K Vs"]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
        unfolding set_add_def' using v by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   556
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
  thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
    using Cons by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   563
subsection \<open>Characterisation of Linearly Independent "Sets"\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   564
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   565
declare independent_backwards [intro]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   566
declare independent_in_carrier [intro]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   567
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   568
lemma independent_distinct: "independent K Us \<Longrightarrow> distinct Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   569
proof (induct Us rule: list.induct)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   570
  case Nil thus ?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   571
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   572
  case Cons thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   573
    using independent_backwards[OF Cons(2)]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   574
          independent_in_carrier[OF Cons(2)]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   575
          Span_base_incl
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   576
    by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   577
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   578
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   579
lemma independent_strinct_incl:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   580
  assumes "independent K (u # Us)" shows "Span K Us \<subset> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   581
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   582
  have "u \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   583
    using Span_base_incl[OF independent_in_carrier[OF assms]] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   584
  moreover have "Span K Us \<subseteq> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
    using mono_Span independent_in_carrier[OF assms] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
  ultimately show ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
    using independent_backwards(1)[OF assms] by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   588
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   590
corollary independent_replacement:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
  assumes "independent K (u # Us)" and "independent K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
  shows "Span K (u # Us) \<subseteq> Span K Vs \<Longrightarrow> (\<exists>v \<in> set Vs. independent K (v # Us))"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
  assume "Span K (u # Us) \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
  hence "Span K Us \<subset> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   596
    using independent_strinct_incl[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
  then obtain v where v: "v \<in> set Vs" "v \<notin> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
    using Span_strict_incl[of Us Vs] assms[THEN independent_in_carrier] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   599
  thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
    using li_Cons[of v K Us] assms independent_in_carrier[OF assms(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
lemma independent_split:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
  assumes "independent K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
  shows "independent K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
    and "independent K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
    and "Span K Us \<inter> Span K Vs = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
  from assms show "independent K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   610
    by (induct Us) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   611
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
  from assms show "independent K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   613
  proof (induct Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   614
    case Nil thus ?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   615
  next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   616
    case (Cons u Us')
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
    hence u: "u \<in> carrier R" and "set Us' \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   618
      using independent_in_carrier[of K "(u # Us') @ Vs"] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
    hence "Span K Us' \<subseteq> Span K (Us' @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   620
      using mono_Span_append(1) by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   621
    thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
      using independent_backwards[of K u "Us' @ Vs"] Cons li_Cons[OF u] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
  from assms show "Span K Us \<inter> Span K Vs = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
  proof (induct Us rule: list.induct)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   627
    case Nil thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
      using Span_subgroup_props(2)[OF independent_in_carrier[of K Vs]] by simp 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
  next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
    case (Cons u Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   631
    hence IH: "Span K Us \<inter> Span K Vs = {\<zero>}" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   632
    have in_carrier:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
      "u \<in> carrier R" "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R" "set (u # Us) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   634
      using Cons(2)[THEN independent_in_carrier] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
    hence "{ \<zero> } \<subseteq> Span K (u # Us) \<inter> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   636
      using in_carrier(3-4)[THEN Span_subgroup_props(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   637
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
    moreover have "Span K (u # Us) \<inter> Span K Vs \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
    proof (rule ccontr)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
      assume "\<not> Span K (u # Us) \<inter> Span K Vs \<subseteq> {\<zero>}"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
      hence "\<exists>a. a \<noteq> \<zero> \<and> a \<in> Span K (u # Us) \<and> a \<in> Span K Vs" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
      then obtain k u' v'
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   643
        where u': "u' \<in> Span K Us" "u' \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
          and v': "v' \<in> Span K Vs" "v' \<in> carrier R" "v' \<noteq> \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
          and k: "k \<in> K" "(k \<otimes> u \<oplus> u') = v'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
        using line_extension_mem_iff[of _ u "Span K Us"] in_carrier(2-3)[THEN Span_subgroup_props(1)]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
              subring_props(1) by force
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
      hence "v' = \<zero>" if "k = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
        using in_carrier(1) that IH by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
      hence diff_zero: "k \<noteq> \<zero>" using v'(3) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
      have "k \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
        using subring_props(1) k(1) by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   654
      hence "k \<otimes> u = (\<ominus> u') \<oplus> v'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
        using in_carrier(1) k(2) u'(2) v'(2) add.m_comm r_neg1 by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
      hence "k \<otimes> u \<in> Span K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
        using Span_subgroup_props(4)[OF in_carrier(2) u'(1)] v'(1) 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
              Span_append_eq_set_add[OF in_carrier(2-3)] unfolding set_add_def' by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
      hence "u \<in> Span K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
        using Cons(2) Span_m_inv_simprule[OF _ _ in_carrier(1), of "Us @ Vs" k]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
              diff_zero k(1) in_carrier(2-3) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
      moreover have "u \<notin> Span K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
        using independent_backwards(1)[of K u "Us @ Vs"] Cons(2) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
      ultimately show False by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
    ultimately show ?case by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
lemma independent_append:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
  assumes "independent K Us" and "independent K Vs" and "Span K Us \<inter> Span K Vs = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
  shows "independent K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
proof (induct Us rule: list.induct)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
  case Nil thus ?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
  case (Cons u Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
  hence in_carrier:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
    "u \<in> carrier R" "set Us \<subseteq> carrier R" "set Vs \<subseteq> carrier R" "set (u # Us) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
    using Cons(2-3)[THEN independent_in_carrier] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
  hence "Span K Us \<subseteq> Span K (u # Us)" 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
    using mono_Span by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   684
  hence "Span K Us \<inter> Span K Vs = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   685
    using Cons(4) Span_subgroup_props(2)[OF in_carrier(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
  hence "independent K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
    using Cons by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
  moreover have "u \<notin> Span K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
  proof (rule ccontr)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
    assume "\<not> u \<notin> Span K (Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
    then obtain u' v'
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   692
      where u': "u' \<in> Span K Us" "u' \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
        and v': "v' \<in> Span K Vs" "v' \<in> carrier R" and u:"u = u' \<oplus> v'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
      using Span_append_eq_set_add[OF in_carrier(2-3)] in_carrier(2-3)[THEN Span_subgroup_props(1)]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
      unfolding set_add_def' by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
    hence "u \<oplus> (\<ominus> u') = v'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
      using in_carrier(1) by algebra
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
    moreover have "u \<in> Span K (u # Us)" and "u' \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
      using Span_base_incl[OF in_carrier(4)] mono_Span[OF in_carrier(2,1)] u'(1)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   700
      by (auto simp del: Span.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
    hence "u \<oplus> (\<ominus> u') \<in> Span K (u # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
      using Span_subgroup_props(3-4)[OF in_carrier(4)] by (auto simp del: Span.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
    ultimately have "u \<oplus> (\<ominus> u') = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
      using Cons(4) v'(1) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   705
    hence "u = u'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
      using Cons(4) v'(1) in_carrier(1) u'(2) \<open>u \<oplus> \<ominus> u' = v'\<close> u by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
    thus False
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
      using u'(1) independent_backwards(1)[OF Cons(2)] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
  ultimately show ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
    using in_carrier(1) li_Cons by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
lemma independent_imp_trivial_combine:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
  assumes "independent K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
  shows "\<And>Ks. \<lbrakk> set Ks \<subseteq> K; combine Ks Us = \<zero> \<rbrakk> \<Longrightarrow> set (take (length Us) Ks) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
proof (induct Us rule: list.induct)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
  case Nil thus ?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
  case (Cons u Us) thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   722
  proof (cases "Ks = []")
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   723
    assume "Ks = []" thus ?thesis by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
  next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
    assume "Ks \<noteq> []"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
    then obtain k Ks' where k: "k \<in> K" and Ks': "set Ks' \<subseteq> K" and Ks: "Ks = k # Ks'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
      using Cons(2) by (metis insert_subset list.exhaust_sel list.simps(15))
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
    hence Us: "set Us \<subseteq> carrier R" and u: "u \<in> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
      using independent_in_carrier[OF Cons(4)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
    have "u \<in> Span K Us" if "k \<noteq> \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
      using that Span_mem_iff[OF Us u] Cons(3-4) Ks' k unfolding Ks by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   732
    hence k_zero: "k = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   733
      using independent_backwards[OF Cons(4)] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
    hence "combine Ks' Us = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
      using combine_in_carrier[OF _ Us, of Ks'] Ks' u Cons(3) subring_props(1) unfolding Ks by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
    hence "set (take (length Us) Ks') \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   737
      using Cons(1)[OF Ks' _ independent_backwards(2)[OF Cons(4)]] by simp 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   738
    thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
      using k_zero unfolding Ks by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
lemma trivial_combine_imp_independent:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
  assumes "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
    and "\<And>Ks. \<lbrakk> set Ks \<subseteq> K; combine Ks Us = \<zero> \<rbrakk> \<Longrightarrow> set (take (length Us) Ks) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
  shows "independent K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
proof (induct Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
  case Nil thus ?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   751
  case (Cons u Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
  hence Us: "set Us \<subseteq> carrier R" and u: "u \<in> carrier R" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
  have "\<And>Ks. \<lbrakk> set Ks \<subseteq> K; combine Ks Us = \<zero> \<rbrakk> \<Longrightarrow> set (take (length Us) Ks) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
  proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
    fix Ks assume Ks: "set Ks \<subseteq> K" and lin_c: "combine Ks Us = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
    hence "combine (\<zero> # Ks) (u # Us) = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
      using u subring_props(1) combine_in_carrier[OF _ Us] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
    hence "set (take (length (u # Us)) (\<zero> # Ks)) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
      using Cons(3)[of "\<zero> # Ks"] subring_props(2) Ks by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
    thus "set (take (length Us) Ks) \<subseteq> { \<zero> }" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
  hence "independent K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
    using Cons(1)[OF Us] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
  moreover have "u \<notin> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
  proof (rule ccontr)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
    assume "\<not> u \<notin> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
    then obtain k Ks where k: "k \<in> K" "k \<noteq> \<zero>" and Ks: "set Ks \<subseteq> K" and u: "combine (k # Ks) (u # Us) = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
      using Span_mem_iff[OF Us u] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   771
    have "set (take (length (u # Us)) (k # Ks)) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   772
      using Cons(3)[OF _ u] k(1) Ks by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
    hence "k = \<zero>" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
    from \<open>k = \<zero>\<close> and \<open>k \<noteq> \<zero>\<close> show False by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
  ultimately show ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
    using li_Cons[OF u] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
corollary unique_decomposition:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
  assumes "independent K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   783
  shows "a \<in> Span K Us \<Longrightarrow> \<exists>!Ks. set Ks \<subseteq> K \<and> length Ks = length Us \<and> a = combine Ks Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
  note in_carrier = independent_in_carrier[OF assms]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
  assume "a \<in> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
  then obtain Ks where Ks: "set Ks \<subseteq> K" "length Ks = length Us" "a = combine Ks Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
    using Span_mem_iff_length_version[OF in_carrier] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
  moreover
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
  have "\<And>Ks'. \<lbrakk> set Ks' \<subseteq> K; length Ks' = length Us; a = combine Ks' Us \<rbrakk> \<Longrightarrow> Ks = Ks'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
  proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   794
    fix Ks' assume Ks': "set Ks' \<subseteq> K" "length Ks' = length Us" "a = combine Ks' Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   795
    hence set_Ks: "set Ks \<subseteq> carrier R" and set_Ks': "set Ks' \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
      using subring_props(1) Ks(1) by blast+
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
    have same_length: "length Ks = length Ks'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
      using Ks Ks' by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
    have "(combine Ks Us) \<oplus> ((\<ominus> \<one>) \<otimes> (combine Ks' Us)) = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
      using combine_in_carrier[OF set_Ks  in_carrier]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   802
            combine_in_carrier[OF set_Ks' in_carrier] Ks(3) Ks'(3) by algebra
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   803
    hence "(combine Ks Us) \<oplus> (combine (map ((\<otimes>) (\<ominus> \<one>)) Ks') Us) = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
      using combine_r_distr[OF set_Ks' in_carrier, of "\<ominus> \<one>"] subring_props by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
    moreover have set_map: "set (map ((\<otimes>) (\<ominus> \<one>)) Ks') \<subseteq> K"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
      using Ks'(1) subring_props by (induct Ks') (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
    hence "set (map ((\<otimes>) (\<ominus> \<one>)) Ks') \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
      using subring_props(1) by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
    ultimately have "combine (map2 (\<oplus>) Ks (map ((\<otimes>) (\<ominus> \<one>)) Ks')) Us = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
      using combine_add[OF Ks(2) _ set_Ks _ in_carrier, of "map ((\<otimes>) (\<ominus> \<one>)) Ks'"] Ks'(2) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
    moreover have "set (map2 (\<oplus>) Ks (map ((\<otimes>) (\<ominus> \<one>)) Ks')) \<subseteq> K"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
      using Ks(1) set_map subring_props(7)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
      by (induct Ks) (auto, metis contra_subsetD in_set_zipE local.set_map set_ConsD subring_props(7))
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
    ultimately have "set (take (length Us) (map2 (\<oplus>) Ks (map ((\<otimes>) (\<ominus> \<one>)) Ks'))) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
      using independent_imp_trivial_combine[OF assms] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
    hence "set (map2 (\<oplus>) Ks (map ((\<otimes>) (\<ominus> \<one>)) Ks')) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
      using Ks(2) Ks'(2) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
    thus "Ks = Ks'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
      using set_Ks set_Ks' same_length
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
    proof (induct Ks arbitrary: Ks')
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
      case Nil thus?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
    next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
      case (Cons k Ks)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
      then obtain k' Ks'' where k': "Ks' = k' # Ks''"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
        by (metis Suc_length_conv)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
      have "Ks = Ks''"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   827
        using Cons unfolding k' by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
      moreover have "k = k'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
        using Cons(2-4) l_minus minus_equality unfolding k' by (auto, fastforce)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
      ultimately show ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
        unfolding k' by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
    qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
  ultimately show ?thesis by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
subsection \<open>Replacement Theorem\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
lemma independent_rotate1_aux:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   842
  "independent K (u # Us @ Vs) \<Longrightarrow> independent K ((Us @ [u]) @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
  assume "independent K (u # Us @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
  hence li: "independent K [u]" "independent K Us" "independent K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
    and inter: "Span K [u] \<inter> Span K Us = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
               "Span K (u # Us) \<inter> Span K Vs = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
    using independent_split[of "u # Us" Vs] independent_split[of "[u]" Us] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
  hence "independent K (Us @ [u])"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
    using independent_append[OF li(2,1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
  moreover have "Span K (Us @ [u]) \<inter> Span K Vs = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
    using Span_same_set[of "u # Us" "Us @ [u]"] li(1-2)[THEN independent_in_carrier] inter(2) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
  ultimately show "independent K ((Us @ [u]) @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
    using independent_append[OF _ li(3), of "Us @ [u]"] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   855
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   857
corollary independent_rotate1:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
  "independent K (Us @ Vs) \<Longrightarrow> independent K ((rotate1 Us) @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   859
  using independent_rotate1_aux by (cases Us) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   860
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
(*
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
corollary independent_rotate:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   863
  "independent K (Us @ Vs) \<Longrightarrow> independent K ((rotate n Us) @ Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
  using independent_rotate1 by (induct n) auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
lemma rotate_append: "rotate (length l) (l @ q) = q @ l"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
  by (induct l arbitrary: q) (auto simp add: rotate1_rotate_swap)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
*)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
corollary independent_same_set:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   871
  assumes "set Us = set Vs" and "length Us = length Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
  shows "independent K Us \<Longrightarrow> independent K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   873
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   874
  assume "independent K Us" thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
    using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   876
  proof (induct Us arbitrary: Vs rule: list.induct)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   877
    case Nil thus ?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   878
  next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
    case (Cons u Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
    then obtain Vs' Vs'' where Vs: "Vs = Vs' @ (u # Vs'')"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
      by (metis list.set_intros(1) split_list)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
    
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
    have in_carrier: "u \<in> carrier R" "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   884
      using independent_in_carrier[OF Cons(2)] by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
    
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
    have "distinct Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
      using Cons(3-4) independent_distinct[OF Cons(2)]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
      by (metis card_distinct distinct_card)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   889
    hence "u \<notin> set (Vs' @ Vs'')" and "u \<notin> set Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
      using independent_distinct[OF Cons(2)] unfolding Vs by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   891
    hence set_eq: "set Us = set (Vs' @ Vs'')" and "length (Vs' @ Vs'') = length Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
      using Cons(3-4) unfolding Vs by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   893
    hence "independent K (Vs' @ Vs'')"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   894
      using Cons(1)[OF independent_backwards(2)[OF Cons(2)]] unfolding Vs by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   895
    hence "independent K (u # (Vs' @ Vs''))"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   896
      using li_Cons Span_same_set[OF _ set_eq] independent_backwards(1)[OF Cons(2)] in_carrier by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
    hence "independent K (Vs' @ (u # Vs''))"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
      using independent_rotate1[of "u # Vs'" Vs''] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   899
    thus ?case unfolding Vs .
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
lemma replacement_theorem:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
  assumes "independent K (Us' @ Us)" and "independent K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
    and "Span K (Us' @ Us) \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
  shows "\<exists>Vs'. set Vs' \<subseteq> set Vs \<and> length Vs' = length Us' \<and> independent K (Vs' @ Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
proof (induct "length Us'" arbitrary: Us' Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
  case 0 thus ?case by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
  case (Suc n)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
  then obtain u Us'' where Us'': "Us' = Us'' @ [u]"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
    by (metis list.size(3) nat.simps(3) rev_exhaust)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
  then obtain Vs' where Vs': "set Vs' \<subseteq> set Vs" "length Vs' = n" "independent K (Vs' @ (u # Us))"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
    using Suc(1)[of Us'' "u # Us"] Suc(2-5) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
  hence li: "independent K ((u # Vs') @ Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
    using independent_same_set[OF _ _ Vs'(3), of "(u # Vs') @ Us"] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
  moreover have in_carrier:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
    "u \<in> carrier R" "set Us \<subseteq> carrier R" "set Us' \<subseteq> carrier R" "set Vs \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   920
    using Suc(3-4)[THEN independent_in_carrier] Us'' by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
  moreover have "Span K ((u # Vs') @ Us) \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   922
  proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
    have "set Us \<subseteq> Span K Vs" "u \<in> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
      using Suc(5) Span_base_incl[of "Us' @ Us"] Us'' in_carrier(2-3) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
    moreover have "set Vs' \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   926
      using Span_base_incl[OF in_carrier(4)] Vs'(1) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
    ultimately have "set ((u # Vs') @ Us) \<subseteq> Span K Vs" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
    thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
      using mono_Span_subset[OF _ in_carrier(4)] by (simp del: Span.simps)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
  ultimately obtain v where "v \<in> set Vs" "independent K ((v # Vs') @ Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
    using independent_replacement[OF _ Suc(4), of u "Vs' @ Us"] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
  thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
    using Vs'(1-2) Suc(2)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
    by (metis (mono_tags, lifting) insert_subset length_Cons list.simps(15))
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   937
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   938
corollary independent_length_le:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   939
  assumes "independent K Us" and "independent K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
  shows "set Us \<subseteq> Span K Vs \<Longrightarrow> length Us \<le> length Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
  assume "set Us \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
  hence "Span K Us \<subseteq> Span K Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
    using mono_Span_subset[OF _ independent_in_carrier[OF assms(2)]] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
  then obtain Vs' where Vs': "set Vs' \<subseteq> set Vs" "length Vs' = length Us" "independent K Vs'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
    using replacement_theorem[OF _ assms(2), of Us "[]"] assms(1) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
  hence "card (set Vs') \<le> card (set Vs)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
    by (simp add: card_mono)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
  thus "length Us \<le> length Vs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
    using independent_distinct assms(2) Vs'(2-3) by (simp add: distinct_card)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   952
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
subsection \<open>Dimension\<close>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   956
lemma exists_base:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
  assumes "dimension n K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
  shows "\<exists>Vs. set Vs \<subseteq> carrier R \<and> independent K Vs \<and> length Vs = n \<and> Span K Vs = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
    (is "\<exists>Vs. ?base K Vs E n")
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
  using assms
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
proof (induct E rule: dimension.induct)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
  case zero_dim thus ?case by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
  case (Suc_dim v E n K)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
  then obtain Vs where Vs: "set Vs \<subseteq> carrier R" "independent K Vs" "length Vs = n" "Span K Vs = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
    by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
  hence "?base K (v # Vs) (line_extension K v E) (Suc n)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
    using Suc_dim li_Cons by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
  thus ?case by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
lemma dimension_zero [intro]: "dimension 0 K E \<Longrightarrow> E = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
  assume "dimension 0 K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   975
  then obtain Vs where "length Vs = 0" "Span K Vs = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   976
    using exists_base by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
  thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
    by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   979
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   980
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   981
lemma dimension_independent [intro]: "independent K Us \<Longrightarrow> dimension (length Us) K (Span K Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   982
proof (induct Us)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   983
  case Nil thus ?case by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   984
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   985
  case Cons thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   986
    using Suc_dim[OF independent_backwards(3,1)[OF Cons(2)]] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   987
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   988
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   989
lemma dimensionI:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
  assumes "independent K Us" "Span K Us = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
  shows "dimension (length Us) K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
  using dimension_independent[OF assms(1)] assms(2) by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   993
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   994
lemma space_subgroup_props:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
  assumes "dimension n K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
  shows "E \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   997
    and "\<zero> \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
    and "\<And>v1 v2. \<lbrakk> v1 \<in> E; v2 \<in> E \<rbrakk> \<Longrightarrow> (v1 \<oplus> v2) \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   999
    and "\<And>v. v \<in> E \<Longrightarrow> (\<ominus> v) \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1000
    and "\<And>k v. \<lbrakk> k \<in> K; v \<in> E \<rbrakk> \<Longrightarrow> k \<otimes> v \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1001
    and "\<lbrakk> k \<in> K - { \<zero> }; a \<in> carrier R \<rbrakk> \<Longrightarrow> k \<otimes> a \<in> E \<Longrightarrow> a \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1002
  using exists_base[OF assms] Span_subgroup_props Span_smult_closed Span_m_inv_simprule by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1003
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1004
lemma independent_length_le_dimension:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1005
  assumes "dimension n K E" and "independent K Us" "set Us \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
  shows "length Us \<le> n"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
  obtain Vs where Vs: "set Vs \<subseteq> carrier R" "independent K Vs" "length Vs = n" "Span K Vs = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
    using exists_base[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1010
  thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
    using independent_length_le assms(2-3) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
lemma dimension_is_inj:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
  assumes "dimension n K E" and "dimension m K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
  shows "n = m"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
  { fix n m assume n: "dimension n K E" and m: "dimension m K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1019
    then obtain Vs
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
      where Vs: "set Vs \<subseteq> carrier R" "independent K Vs" "length Vs = n" "Span K Vs = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
      using exists_base by meson
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
    hence "n \<le> m"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
      using independent_length_le_dimension[OF m Vs(2)] Span_base_incl[OF Vs(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
  } note aux_lemma = this
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1026
  show ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
    using aux_lemma[OF assms] aux_lemma[OF assms(2,1)] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1028
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
corollary independent_length_eq_dimension:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
  assumes "dimension n K E" and "independent K Us" "set Us \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1032
  shows "length Us = n \<longleftrightarrow> Span K Us = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1034
  assume len: "length Us = n" show "Span K Us = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1035
  proof (rule ccontr)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1036
    assume "Span K Us \<noteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1037
    hence "Span K Us \<subset> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1038
      using mono_Span_subset[of Us] exists_base[OF assms(1)] assms(3) by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1039
    then obtain v where v: "v \<in> E" "v \<notin> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1040
      using Span_strict_incl exists_base[OF assms(1)] space_subgroup_props(1)[OF assms(1)] assms by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1041
    hence "independent K (v # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1042
      using li_Cons[OF _ _ assms(2)] space_subgroup_props(1)[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1043
    hence "length (v # Us) \<le> n"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1044
      using independent_length_le_dimension[OF assms(1)] v(1) assms(2-3) by fastforce
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
    moreover have "length (v # Us) = Suc n"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
      using len by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
    ultimately show False by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1048
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
  assume "Span K Us = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
  hence "dimension (length Us) K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1052
    using dimensionI assms by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
  thus "length Us = n"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
    using dimension_is_inj[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
lemma complete_base:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
  assumes "dimension n K E" and "independent K Us" "set Us \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
  shows "\<exists>Vs. length (Vs @ Us) = n \<and> independent K (Vs @ Us) \<and> Span K (Vs @ Us) = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
  { fix Us k assume "k \<le> n" "independent K Us" "set Us \<subseteq> E" "length Us = k"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
    hence "\<exists>Vs. length (Vs @ Us) = n \<and> independent K (Vs @ Us) \<and> Span K (Vs @ Us) = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
    proof (induct arbitrary: Us rule: inc_induct)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
      case base thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
        using independent_length_eq_dimension[OF assms(1) base(1-2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
    next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
      case (step m)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
      have "Span K Us \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
        using mono_Span_subset step(4-6) exists_base[OF assms(1)] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
      hence "Span K Us \<subset> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
        using independent_length_eq_dimension[OF assms(1) step(4-5)] step(2,6) assms(1) by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
      then obtain v where v: "v \<in> E" "v \<notin> Span K Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
        using Span_strict_incl exists_base[OF assms(1)] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
      hence "independent K (v # Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
        using space_subgroup_props(1)[OF assms(1)] li_Cons[OF _ v(2) step(4)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
      then obtain Vs
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
        where "length (Vs @ (v # Us)) = n" "independent K (Vs @ (v # Us))" "Span K (Vs @ (v # Us)) = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
        using step(3)[of "v # Us"] step(1-2,4-6) v by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1079
      thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
        by (metis append.assoc append_Cons append_Nil)  
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
    qed } note aux_lemma = this
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
  have "length Us \<le> n"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
    using independent_length_le_dimension[OF assms] .
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
  thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
    using aux_lemma[OF _ assms(2-3)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
lemma dimension_backwards:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1090
  "dimension (Suc n) K E \<Longrightarrow> \<exists>v \<in> carrier R. \<exists>E'. dimension n K E' \<and> v \<notin> E' \<and> E = line_extension K v E'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
  by (cases rule: dimension.cases) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
lemma dimension_direct_sum_space:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1094
  assumes "dimension n K E" and "dimension m K F" and "E \<inter> F = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1095
  shows "dimension (n + m) K (E <+>\<^bsub>R\<^esub> F)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
  obtain Us Vs
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1098
    where Vs: "set Vs \<subseteq> carrier R" "independent K Vs" "length Vs = n" "Span K Vs = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1099
      and Us: "set Us \<subseteq> carrier R" "independent K Us" "length Us = m" "Span K Us = F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1100
    using assms(1-2)[THEN exists_base] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1101
  hence "Span K (Vs @ Us) = E <+>\<^bsub>R\<^esub> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
    using Span_append_eq_set_add by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
  moreover have "independent K (Vs @ Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
    using assms(3) independent_append[OF Vs(2) Us(2)] unfolding Vs(4) Us(4) by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1105
  ultimately show "dimension (n + m) K (E <+>\<^bsub>R\<^esub> F)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
    using dimensionI[of "Vs @ Us"] Vs(3) Us(3) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
lemma dimension_sum_space:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
  assumes "dimension n K E" and "dimension m K F" and "dimension k K (E \<inter> F)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
  shows "dimension (n + m - k) K (E <+>\<^bsub>R\<^esub> F)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1112
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
  obtain Bs
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1114
    where Bs: "set Bs \<subseteq> carrier R" "length Bs = k" "independent K Bs" "Span K Bs = E \<inter> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1115
    using exists_base[OF assms(3)] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1116
  then obtain Us Vs
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1117
    where Us: "length (Us @ Bs) = n" "independent K (Us @ Bs)" "Span K (Us @ Bs) = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1118
      and Vs: "length (Vs @ Bs) = m" "independent K (Vs @ Bs)" "Span K (Vs @ Bs) = F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1119
    using Span_base_incl[OF Bs(1)] assms(1-2)[THEN complete_base] by (metis le_infE)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1120
  hence in_carrier: "set Us \<subseteq> carrier R" "set (Vs @ Bs) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1121
    using independent_in_carrier[OF Us(2)] independent_in_carrier[OF Vs(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1122
  hence "Span K Us \<inter> (Span K (Vs @ Bs)) \<subseteq> Span K Bs"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
    using Bs(4) Us(3) Vs(3) mono_Span_append(1)[OF _ Bs(1), of Us] by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1124
  hence "Span K Us \<inter> (Span K (Vs @ Bs)) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
    using independent_split(3)[OF Us(2)] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
  hence "Span K Us \<inter> (Span K (Vs @ Bs)) = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
    using in_carrier[THEN Span_subgroup_props(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1128
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
  hence dim: "dimension (n + m - k) K (Span K (Us @ (Vs @ Bs)))"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
    using independent_append[OF independent_split(2)[OF Us(2)] Vs(2)] Us(1) Vs(1) Bs(2)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
          dimension_independent[of "Us @ (Vs @ Bs)"] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
  have "(Span K Us) <+>\<^bsub>R\<^esub> F \<subseteq> E <+>\<^bsub>R\<^esub> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
    using mono_Span_append(1)[OF in_carrier(1) Bs(1)] Us(3) unfolding set_add_def' by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
  moreover have "E <+>\<^bsub>R\<^esub> F \<subseteq> (Span K Us) <+>\<^bsub>R\<^esub> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1136
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
    fix v assume "v \<in> E <+>\<^bsub>R\<^esub> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1138
    then obtain u' v' where v: "u' \<in> E" "v' \<in> F" "v = u' \<oplus> v'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1139
      unfolding set_add_def' by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
    then obtain u1' u2' where u1': "u1' \<in> Span K Us" and u2': "u2' \<in> Span K Bs" and u': "u' = u1' \<oplus> u2'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
      using Span_append_eq_set_add[OF in_carrier(1) Bs(1)] Us(3) unfolding set_add_def' by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
    have "v = u1' \<oplus> (u2' \<oplus> v')"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
      using Span_subgroup_props(1)[OF Bs(1)] Span_subgroup_props(1)[OF in_carrier(1)]
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
            space_subgroup_props(1)[OF assms(2)] u' v u1' u2' a_assoc[of u1' u2' v'] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
    moreover have "u2' \<oplus> v' \<in> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
      using space_subgroup_props(3)[OF assms(2) _ v(2)] u2' Bs(4) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1148
    ultimately show "v \<in> (Span K Us) <+>\<^bsub>R\<^esub> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1149
      using u1' unfolding set_add_def' by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
  ultimately have "Span K (Us @ (Vs @ Bs)) = E <+>\<^bsub>R\<^esub> F" 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1152
    using Span_append_eq_set_add[OF in_carrier] Vs(3) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
  thus ?thesis using dim by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1155
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
end
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1158
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
end
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
lemma (in ring) telescopic_base_aux:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
  assumes "subfield K R" "subfield F R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
    and "dimension n K F" and "dimension 1 F E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1164
  shows "dimension n K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
  obtain Us u
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
    where Us: "set Us \<subseteq> carrier R" "length Us = n" "independent K Us" "Span K Us = F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1168
      and u: "u \<in> carrier R" "independent F [u]" "Span F [u] = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
    using exists_base[OF assms(2,4)] exists_base[OF assms(1,3)] independent_backwards(3) assms(2)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
    by (metis One_nat_def length_0_conv length_Suc_conv)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
  have in_carrier: "set (map (\<lambda>u'. u' \<otimes> u) Us) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
    using Us(1) u(1) by (induct Us) (auto)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1173
  
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1174
  have li: "independent K (map (\<lambda>u'. u' \<otimes> u) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1175
  proof (rule trivial_combine_imp_independent[OF assms(1) in_carrier])
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1176
    fix Ks assume Ks: "set Ks \<subseteq> K" and "combine Ks (map (\<lambda>u'. u' \<otimes> u) Us) = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
    hence "(combine Ks Us) \<otimes> u = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
      using combine_l_distr[OF _ Us(1) u(1)] subring_props(1)[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1179
    hence "combine [ combine Ks Us ] [ u ] = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1180
      by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1181
    moreover have "combine Ks Us \<in> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1182
      using Us(4) Ks(1) Span_eq_combine_set[OF assms(1) Us(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
    ultimately have "combine Ks Us = \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
      using independent_imp_trivial_combine[OF assms(2) u(2), of "[ combine Ks Us ]"] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
    hence "set (take (length Us) Ks) \<subseteq> { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1186
      using independent_imp_trivial_combine[OF assms(1) Us(3) Ks(1)] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
    thus "set (take (length (map (\<lambda>u'. u' \<otimes> u) Us)) Ks) \<subseteq> { \<zero> }" by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
  have "E \<subseteq> Span K (map (\<lambda>u'. u' \<otimes> u) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
    fix v assume "v \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
    then obtain f where f: "f \<in> F" "v = f \<otimes> u \<oplus> \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
      using u(1,3) line_extension_mem_iff[OF assms(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
    then obtain Ks where Ks: "set Ks \<subseteq> K" "f = combine Ks Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1196
      using Span_eq_combine_set[OF assms(1) Us(1)] Us(4) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
    have "v = f \<otimes> u"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1198
      using subring_props(1)[OF assms(2)] f u(1) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
    hence "v = combine Ks (map (\<lambda>u'. u' \<otimes> u) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1200
      using combine_l_distr[OF _ Us(1) u(1), of Ks] Ks(1-2)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
            subring_props(1)[OF assms(1)] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
    thus "v \<in> Span K (map (\<lambda>u'. u' \<otimes> u) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
      unfolding Span_eq_combine_set[OF assms(1) in_carrier] using Ks(1) by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1205
  moreover have "Span K (map (\<lambda>u'. u' \<otimes> u) Us) \<subseteq> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
  proof
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
    fix v assume "v \<in> Span K (map (\<lambda>u'. u' \<otimes> u) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
    then obtain Ks where Ks: "set Ks \<subseteq> K" "v = combine Ks (map (\<lambda>u'. u' \<otimes> u) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
      unfolding Span_eq_combine_set[OF assms(1) in_carrier] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
    hence "v = (combine Ks Us) \<otimes> u"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
      using combine_l_distr[OF _ Us(1) u(1), of Ks] subring_props(1)[OF assms(1)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
    moreover have "combine Ks Us \<in> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
      using Us(4) Span_eq_combine_set[OF assms(1) Us(1)] Ks(1) by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
    ultimately have "v = (combine Ks Us) \<otimes> u \<oplus> \<zero>" and "combine Ks Us \<in> F"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
      using subring_props(1)[OF assms(2)] u(1) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
    thus "v \<in> E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
      using u(3) line_extension_mem_iff[OF assms(2)] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
  qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
  ultimately have "Span K (map (\<lambda>u'. u' \<otimes> u) Us) = E" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
  thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
    using dimensionI[OF assms(1) li] Us(2) by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
lemma (in ring) telescopic_base:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
  assumes "subfield K R" "subfield F R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1226
    and "dimension n K F" and "dimension m F E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1227
  shows "dimension (n * m) K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1228
  using assms(4)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1229
proof (induct m arbitrary: E)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
  case 0 thus ?case
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
    using dimension_zero[OF assms(2)] zero_dim by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
  case (Suc m)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
  obtain Vs
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1235
    where Vs: "set Vs \<subseteq> carrier R" "length Vs = Suc m" "independent F Vs" "Span F Vs = E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
    using exists_base[OF assms(2) Suc(2)] by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
  then obtain v Vs' where v: "Vs = v # Vs'"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
    by (meson length_Suc_conv)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
  hence li: "independent F [ v ]" "independent F Vs'" and inter: "Span F [ v ] \<inter> Span F Vs' = { \<zero> }"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1240
    using Vs(3) independent_split[OF assms(2), of "[ v ]" Vs'] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
  have "dimension n K (Span F [ v ])"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
    using dimension_independent[OF assms(2) li(1)] telescopic_base_aux[OF assms(1-3)] by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1243
  moreover have "dimension (n * m) K (Span F Vs')"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1244
    using Suc(1) dimension_independent[OF assms(2) li(2)] Vs(2) unfolding v by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
  ultimately have "dimension (n * Suc m) K (Span F [ v ] <+>\<^bsub>R\<^esub> Span F Vs')"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
    using dimension_direct_sum_space[OF assms(1) _ _ inter] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
  thus "dimension (n * Suc m) K E"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1248
    using Span_append_eq_set_add[OF assms(2) li[THEN independent_in_carrier]] Vs(4) v by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1250
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
(*
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1253
lemma combine_take:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
  assumes "set Ks  \<subseteq> carrier R" "set Us \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
  shows "length Ks \<le> length Us \<Longrightarrow> combine Ks Us = combine Ks (take (length Ks) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
    and "length Us \<le> length Ks \<Longrightarrow> combine Ks Us = combine (take (length Us) Ks) Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
proof -
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
  assume len: "length Ks \<le> length Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
  hence Us: "Us = (take (length Ks) Us) @ (drop (length Ks) Us)" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
  hence set_t: "set (take (length Ks) Us) \<subseteq> carrier R" and set_d: "set (drop (length Ks) Us) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
    using assms(2) len by (metis le_sup_iff set_append)+
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
  hence "combine Ks Us = (combine Ks (take (length Ks) Us)) \<oplus> \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
    using combine_append[OF _ assms(1), of "take (length Ks) Us" "[]" "drop (length Ks) Us"] len by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
  also have " ... = combine Ks (take (length Ks) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
    using combine_in_carrier[OF assms(1) set_t] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1266
  finally show "combine Ks Us = combine Ks (take (length Ks) Us)" .
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1267
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1268
  assume len: "length Us \<le> length Ks"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1269
  hence Us: "Ks = (take (length Us) Ks) @ (drop (length Us) Ks)" by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
  hence set_t: "set (take (length Us) Ks) \<subseteq> carrier R" and set_d: "set (drop (length Us) Ks) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
    using assms(1) len by (metis le_sup_iff set_append)+
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1272
  hence "combine Ks Us = (combine (take (length Us) Ks) Us) \<oplus> \<zero>"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1273
    using combine_append[OF _ _ assms(2), of "take (length Us) Ks" "drop (length Us) Ks" "[]"] len by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1274
  also have " ... = combine (take (length Us) Ks) Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1275
    using combine_in_carrier[OF set_t assms(2)] by auto 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1276
  finally show "combine Ks Us = combine (take (length Us) Ks) Us" .
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1277
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1278
*)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1279
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1280
(*
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1281
lemma combine_normalize:
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1282
  assumes "set Ks \<subseteq> K" "set Us \<subseteq> carrier R" "a = combine Ks Us" 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1283
  shows "\<exists>Ks'. set Ks' \<subseteq> K \<and> length Ks' = length Us \<and> a = combine Ks' Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1284
proof (cases "length Ks \<le> length Us")
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1285
  assume "\<not> length Ks \<le> length Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1286
  hence len: "length Us < length Ks" by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
  hence "length (take (length Us) Ks) = length Us" and "set (take (length Us) Ks) \<subseteq> K"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
    using assms(1) by (auto, metis contra_subsetD in_set_takeD)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
  thus ?thesis
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
    using combine_take(2)[OF _ assms(2), of Ks] assms(1,3) subring_props(1) len
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
    by (metis dual_order.trans nat_less_le)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1292
next
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
  assume len: "length Ks \<le> length Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
  have Ks: "set Ks \<subseteq> carrier R" and set_r: "set (replicate (length Us - length Ks) \<zero>) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
    using assms subring_props(1) zero_closed by (metis dual_order.trans, auto) 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
  moreover
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
  have set_t: "set (take (length Ks) Us) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
   and set_d: "set (drop (length Ks) Us) \<subseteq> carrier R"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
    using assms(2) len dual_order.trans by (metis set_take_subset, metis set_drop_subset)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1300
  ultimately 
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
  have "combine (Ks @ (replicate (length Us - length Ks) \<zero>)) Us =
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
       (combine Ks (take (length Ks) Us)) \<oplus>
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
       (combine (replicate (length Us - length Ks) \<zero>) (drop (length Ks) Us))"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
    using combine_append[OF _ Ks set_t set_r set_d] len by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
  also have " ... = combine Ks (take (length Ks) Us)"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
    using combine_replicate[OF set_d] combine_in_carrier[OF Ks set_t] by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
  also have " ... = a"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
    using combine_take(1)[OF Ks assms(2) len] assms(3) by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
  finally have "combine (Ks @ (replicate (length Us - length Ks) \<zero>)) Us = a" .
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
  moreover have "set (Ks @ (replicate (length Us - length Ks) \<zero>)) \<subseteq> K"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
    using assms(1) subring_props(2) by auto
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
  moreover have "length (Ks @ (replicate (length Us - length Ks) \<zero>)) = length Us"
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
    using len by simp
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
  ultimately show ?thesis by blast
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
qed
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
*)
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
c64319959bab Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
end