author | wenzelm |
Sat, 13 Aug 2022 18:06:30 +0200 | |
changeset 75848 | 9e4c0aaa30aa |
parent 70160 | 8e9100dcde52 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
68582 | 1 |
(* Title: HOL/Algebra/Embedded_Algebras.thy |
2 |
Author: Paulo Emílio de Vilhena |
|
3 |
*) |
|
68569
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 |
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
|
6 |
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
|
7 |
begin |
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 |
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
|
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 |
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
|
12 |
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
|
13 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
14 |
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
|
15 |
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
|
16 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
17 |
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
|
18 |
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
|
19 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
20 |
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
|
21 |
where |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
22 |
"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
|
23 |
| "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
|
24 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
25 |
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
|
26 |
where |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
27 |
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
|
28 |
| 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
|
29 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
30 |
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
|
31 |
where |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
32 |
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
|
33 |
| 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
|
34 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
35 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
36 |
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
|
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 |
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
|
39 |
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
|
40 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
41 |
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
|
42 |
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
|
43 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
44 |
|
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 |
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
|
47 |
begin |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
48 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
49 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
50 |
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
|
51 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
52 |
lemma line_extension_consistent: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
53 |
assumes "subring K R" shows "ring.line_extension (R \<lparr> carrier := K \<rparr>) = line_extension" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
54 |
unfolding ring.line_extension_def[OF subring_is_ring[OF assms]] line_extension_def |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
55 |
by (simp add: set_add_def set_mult_def) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
56 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
57 |
lemma Span_consistent: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
58 |
assumes "subring K R" shows "ring.Span (R \<lparr> carrier := K \<rparr>) = Span" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
59 |
unfolding ring.Span.simps[OF subring_is_ring[OF assms]] Span.simps |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
60 |
line_extension_consistent[OF assms] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
61 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
62 |
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
|
63 |
"\<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
|
64 |
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
|
65 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
66 |
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
|
67 |
"\<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
|
68 |
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
|
69 |
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
|
70 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
71 |
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
|
72 |
"\<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
|
73 |
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
|
74 |
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
|
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_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
|
77 |
"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
|
78 |
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
|
79 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
80 |
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
|
81 |
"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
|
82 |
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
|
83 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
84 |
lemma combine_take: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
85 |
"combine (take (length Us) Ks) Us = combine Ks Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
86 |
by (induct Us arbitrary: Ks) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
87 |
(auto, metis combine.simps(1) list.exhaust take.simps(1) take_Suc_Cons) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
88 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
89 |
lemma combine_append_zero: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
90 |
"set Us \<subseteq> carrier R \<Longrightarrow> combine (Ks @ [ \<zero> ]) Us = combine Ks Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
91 |
proof (induct Ks arbitrary: Us) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
92 |
case Nil thus ?case by (induct Us) (auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
93 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
94 |
case Cons thus ?case by (cases Us) (auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
95 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
96 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
97 |
lemma combine_prepend_replicate: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
98 |
"\<lbrakk> set Ks \<subseteq> carrier R; set Us \<subseteq> carrier R \<rbrakk> \<Longrightarrow> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
99 |
combine ((replicate n \<zero>) @ Ks) Us = combine Ks (drop n Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
100 |
proof (induct n arbitrary: Us, simp) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
101 |
case (Suc n) thus ?case |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
102 |
by (cases Us) (auto, meson combine_in_carrier ring_simprules(8) set_drop_subset subset_trans) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
103 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
104 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
105 |
lemma combine_append_replicate: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
106 |
"set Us \<subseteq> carrier R \<Longrightarrow> combine (Ks @ (replicate n \<zero>)) Us = combine Ks Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
107 |
by (induct n) (auto, metis append.assoc combine_append_zero replicate_append_same) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
108 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
109 |
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
|
110 |
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
|
111 |
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
|
112 |
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
|
113 |
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
|
114 |
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
|
115 |
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
|
116 |
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
|
117 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
118 |
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
|
119 |
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
|
120 |
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
|
121 |
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
|
122 |
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
|
123 |
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
|
124 |
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
|
125 |
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
|
126 |
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
|
127 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
128 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
129 |
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
|
130 |
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
|
131 |
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
|
132 |
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
|
133 |
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
|
134 |
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
|
135 |
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
|
136 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
137 |
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
|
138 |
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
|
139 |
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
|
140 |
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
|
141 |
"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
|
142 |
"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
|
143 |
"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
|
144 |
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
|
145 |
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
|
146 |
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
|
147 |
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
|
148 |
((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
|
149 |
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
|
150 |
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
|
151 |
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
|
152 |
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
|
153 |
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
|
154 |
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
|
155 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
156 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
157 |
lemma combine_normalize: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
158 |
assumes "set Ks \<subseteq> carrier R" "set Us \<subseteq> carrier R" "combine Ks Us = a" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
159 |
obtains Ks' |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
160 |
where "set (take (length Us) Ks) \<subseteq> set Ks'" "set Ks' \<subseteq> set (take (length Us) Ks) \<union> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
161 |
and "length Ks' = length Us" "combine Ks' Us = a" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
162 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
163 |
define Ks' |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
164 |
where "Ks' = (if length Ks \<le> length Us |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
165 |
then Ks @ (replicate (length Us - length Ks) \<zero>) else take (length Us) Ks)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
166 |
hence "set (take (length Us) Ks) \<subseteq> set Ks'" "set Ks' \<subseteq> set (take (length Us) Ks) \<union> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
167 |
"length Ks' = length Us" "a = combine Ks' Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
168 |
using combine_append_replicate[OF assms(2)] combine_take assms(3) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
169 |
thus thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
170 |
using that by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
171 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
172 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
173 |
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)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
174 |
unfolding line_extension_def set_add_def'[of R "K #> a" E] unfolding r_coset_def by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
175 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
176 |
lemma line_extension_in_carrier: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
177 |
assumes "K \<subseteq> carrier R" "a \<in> carrier R" "E \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
178 |
shows "line_extension K a E \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
179 |
using set_add_closed[OF r_coset_subset_G[OF assms(1-2)] assms(3)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
180 |
by (simp add: line_extension_def) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
181 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
182 |
lemma Span_in_carrier: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
183 |
assumes "K \<subseteq> carrier R" "set Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
184 |
shows "Span K Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
185 |
using assms by (induct Us) (auto simp add: line_extension_in_carrier) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
186 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
187 |
|
68571 | 188 |
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
|
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 |
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
|
191 |
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
|
192 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
193 |
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
|
194 |
"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
|
195 |
"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
|
196 |
"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
|
197 |
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
|
198 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
199 |
lemma dimension_independent [intro]: "independent K Us \<Longrightarrow> dimension (length Us) K (Span K Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
200 |
proof (induct Us) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
201 |
case Nil thus ?case by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
202 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
203 |
case Cons thus ?case |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
204 |
using Suc_dim independent_backwards[OF Cons(2)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
205 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
206 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
207 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
208 |
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
|
209 |
structures, but our interest is to work with subfields, so generalization could |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
210 |
be the subject of a future work.\<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
|
211 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
212 |
context |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
213 |
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
|
214 |
begin |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
215 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
216 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
217 |
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
|
218 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
219 |
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
|
220 |
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
|
221 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
222 |
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
|
223 |
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
|
224 |
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
|
225 |
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
|
226 |
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
|
227 |
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
|
228 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
229 |
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
|
230 |
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
|
231 |
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
|
232 |
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
|
233 |
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
|
234 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
235 |
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
|
236 |
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
|
237 |
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
|
238 |
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
|
239 |
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
|
240 |
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
|
241 |
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
|
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 |
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
|
244 |
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
|
245 |
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
|
246 |
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
|
247 |
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
|
248 |
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
|
249 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
250 |
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
|
251 |
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
|
252 |
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
|
253 |
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
|
254 |
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
|
255 |
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
|
256 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
257 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
258 |
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
|
259 |
"set Us \<subseteq> carrier R \<Longrightarrow> subgroup (Span K Us) (add_monoid R)" |
68687 | 260 |
using line_extension_is_subgroup normal_imp_subgroup[OF add.one_is_normal] by (induct Us) (auto) |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
261 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
262 |
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
|
263 |
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
|
264 |
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
|
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 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
|
267 |
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
|
268 |
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
|
269 |
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
|
270 |
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
|
271 |
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
|
272 |
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
|
273 |
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
|
274 |
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
|
275 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
276 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
277 |
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
|
278 |
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
|
279 |
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
|
280 |
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
|
281 |
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
|
282 |
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
|
283 |
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
|
284 |
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
|
285 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
286 |
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
|
287 |
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
|
288 |
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
|
289 |
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
|
290 |
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
|
291 |
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
|
292 |
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
|
293 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
294 |
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
|
295 |
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
|
296 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
297 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
298 |
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
|
299 |
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
|
300 |
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
|
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 |
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
|
303 |
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
|
304 |
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
|
305 |
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
|
306 |
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
|
307 |
thus ?thesis |
68604 | 308 |
using inv_k subring_props(1)a k |
309 |
by (metis (no_types, lifting) DiffE l_one m_assoc subset_iff) |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
310 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
311 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
312 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
313 |
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
|
314 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
315 |
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
|
316 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
317 |
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
|
318 |
assumes "u \<in> carrier R" |
68687 | 319 |
shows "line_extension K u { combine Ks Us | Ks. set Ks \<subseteq> K } = |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
320 |
{ 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
|
321 |
(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
|
322 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
323 |
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
|
324 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
325 |
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
|
326 |
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
|
327 |
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
|
328 |
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
|
329 |
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
|
330 |
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
|
331 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
332 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
333 |
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
|
334 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
335 |
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
|
336 |
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
|
337 |
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
|
338 |
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
|
339 |
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
|
340 |
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
|
341 |
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
|
342 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
343 |
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
|
344 |
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
|
345 |
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
|
346 |
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
|
347 |
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
|
348 |
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
|
349 |
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
|
350 |
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
|
351 |
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
|
352 |
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
|
353 |
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
|
354 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
355 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
356 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
357 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
358 |
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
|
359 |
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
|
360 |
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
|
361 |
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
|
362 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
363 |
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
|
364 |
assumes "u \<in> carrier R" |
68687 | 365 |
shows "line_extension K u { combine Ks Us | Ks. length Ks = length Us \<and> set Ks \<subseteq> K } = |
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 |
{ 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
|
367 |
(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
|
368 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
369 |
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
|
370 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
371 |
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
|
372 |
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
|
373 |
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
|
374 |
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
|
375 |
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
|
376 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
377 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
378 |
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
|
379 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
380 |
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
|
381 |
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
|
382 |
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
|
383 |
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
|
384 |
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
|
385 |
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
|
386 |
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
|
387 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
388 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
389 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
390 |
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
|
391 |
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
|
392 |
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
|
393 |
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
|
394 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
395 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
396 |
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
|
397 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
398 |
corollary Span_mem_iff_length_version: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
399 |
assumes "set Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
400 |
shows "a \<in> Span K Us \<longleftrightarrow> (\<exists>Ks. set Ks \<subseteq> K \<and> length Ks = length Us \<and> a = combine Ks Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
401 |
using Span_eq_combine_set_length_version[OF assms] by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
402 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
403 |
corollary Span_mem_imp_non_trivial_combine: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
404 |
assumes "set Us \<subseteq> carrier R" and "a \<in> Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
405 |
obtains k Ks |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
406 |
where "k \<in> K - { \<zero> }" "set Ks \<subseteq> K" "length Ks = length Us" "combine (k # Ks) (a # Us) = \<zero>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
407 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
408 |
obtain Ks where Ks: "set Ks \<subseteq> K" "length Ks = length Us" "a = combine Ks Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
409 |
using Span_mem_iff_length_version[OF assms(1)] assms(2) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
410 |
hence "((\<ominus> \<one>) \<otimes> a) \<oplus> a = combine ((\<ominus> \<one>) # Ks) (a # Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
411 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
412 |
moreover have "((\<ominus> \<one>) \<otimes> a) \<oplus> a = \<zero>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
413 |
using assms(2) Span_subgroup_props(1)[OF assms(1)] l_minus l_neg by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
414 |
moreover have "\<ominus> \<one> \<noteq> \<zero>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
415 |
using subfieldE(6)[OF K] l_neg by force |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
416 |
ultimately show ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
417 |
using that subring_props(3,5) Ks(1-2) by (force simp del: combine.simps) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
418 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
419 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
420 |
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
|
421 |
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
|
422 |
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
|
423 |
(is "?in_Span \<longleftrightarrow> ?exists_combine") |
68687 | 424 |
proof |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
425 |
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
|
426 |
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
|
427 |
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
|
428 |
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
|
429 |
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 |
moreover have "((\<ominus> \<one>) \<otimes> a) \<oplus> a = \<zero>" |
68687 | 431 |
using assms(2) l_minus l_neg by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
432 |
moreover have "\<ominus> \<one> \<noteq> \<zero>" |
68687 | 433 |
using subfieldE(6)[OF K] l_neg by force |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
434 |
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
|
435 |
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
|
436 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
437 |
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
|
438 |
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
|
439 |
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
|
440 |
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
|
441 |
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
|
442 |
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
|
443 |
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
|
444 |
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
|
445 |
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
|
446 |
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
|
447 |
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
|
448 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
449 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
450 |
|
69597 | 451 |
subsection \<open>Span as the minimal subgroup that contains \<^term>\<open>K <#> (set Us)\<close>\<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
|
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 |
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
|
454 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
455 |
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
|
456 |
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
|
457 |
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
|
458 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
459 |
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
|
460 |
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
|
461 |
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
|
462 |
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
|
463 |
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
|
464 |
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
|
465 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
466 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
467 |
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
|
468 |
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
|
469 |
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
|
470 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
471 |
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
|
472 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
473 |
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
|
474 |
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
|
475 |
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
|
476 |
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
|
477 |
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
|
478 |
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
|
479 |
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
|
480 |
hence "k \<in> K" and "u \<in> set (u # Us)" by auto |
68687 | 481 |
hence "k \<otimes> u \<in> E" |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
482 |
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
|
483 |
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
|
484 |
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
|
485 |
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
|
486 |
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
|
487 |
ultimately show ?case |
68687 | 488 |
using add.subgroupE(4)[OF assms(2)] by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
489 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
490 |
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
|
491 |
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
|
492 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
493 |
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
|
494 |
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
|
495 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
496 |
qed |
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_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
|
500 |
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
|
501 |
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
|
502 |
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
|
503 |
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
|
504 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
505 |
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
|
506 |
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
|
507 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
508 |
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
|
509 |
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
|
510 |
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
|
511 |
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
|
512 |
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
|
513 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
514 |
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
|
515 |
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
|
516 |
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
|
517 |
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
|
518 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
519 |
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
|
520 |
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
|
521 |
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
|
522 |
moreover have "k \<in> carrier R" and "u \<in> carrier R" |
68687 | 523 |
using Cons(2) k subring_props(1) by (blast, auto) |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
524 |
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
|
525 |
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
|
526 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
527 |
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
|
528 |
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
|
529 |
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
|
530 |
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
|
531 |
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
|
532 |
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
|
533 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
534 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
535 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
536 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
537 |
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
|
538 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
539 |
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
|
540 |
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
|
541 |
shows "set Us = set Vs \<Longrightarrow> Span K Us = Span K Vs" |
68687 | 542 |
using Span_eq_generate assms by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
543 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
544 |
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
|
545 |
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
|
546 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
547 |
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
|
548 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
549 |
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
|
550 |
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
|
551 |
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
|
552 |
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
|
553 |
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
|
554 |
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
|
555 |
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
|
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 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
558 |
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
|
559 |
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
|
560 |
shows "Span K Us \<subseteq> Span K Vs" |
69122
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents:
68687
diff
changeset
|
561 |
using add.mono_generate[OF mono_set_mult[OF _ assms(1), of K K R]] |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
562 |
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
|
563 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
564 |
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
|
565 |
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
|
566 |
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
|
567 |
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
|
568 |
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
|
569 |
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
|
570 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
571 |
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
|
572 |
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
|
573 |
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
|
574 |
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
|
575 |
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
|
576 |
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
|
577 |
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
|
578 |
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
|
579 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
580 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
581 |
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
|
582 |
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
|
583 |
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
|
584 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
585 |
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
|
586 |
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
|
587 |
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
|
588 |
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
|
589 |
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
|
590 |
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
|
591 |
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
|
592 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
593 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
594 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
595 |
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
|
596 |
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
|
597 |
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
|
598 |
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
|
599 |
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
|
600 |
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
|
601 |
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
|
602 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
603 |
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
|
604 |
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
|
605 |
"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
|
606 |
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
|
607 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
608 |
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
|
609 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
610 |
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
|
611 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
612 |
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
|
613 |
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
|
614 |
where v: "k \<in> K" "u' \<in> Span K Us" "v' \<in> Span K Vs" "v = k \<otimes> u \<oplus> (u' \<oplus> v')" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
615 |
using line_extension_mem_iff[of v _ u "Span K Us <+>\<^bsub>R\<^esub> Span K Vs"] |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
616 |
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
|
617 |
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
|
618 |
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
|
619 |
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
|
620 |
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
|
621 |
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
|
622 |
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
|
623 |
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
|
624 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
625 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
626 |
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
|
627 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
628 |
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
|
629 |
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
|
630 |
where v: "k \<in> K" "u' \<in> Span K Us" "v' \<in> Span K Vs" "v = (k \<otimes> u \<oplus> u') \<oplus> v'" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
631 |
using line_extension_mem_iff[of _ _ u "Span K Us"] unfolding set_add_def' by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
632 |
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
|
633 |
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
|
634 |
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
|
635 |
thus "v \<in> line_extension K u (Span K Us <+>\<^bsub>R\<^esub> Span K Vs)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
636 |
using line_extension_mem_iff[of "(k \<otimes> u) \<oplus> (u' \<oplus> v')" K u "Span K Us <+>\<^bsub>R\<^esub> Span K Vs"] |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
637 |
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
|
638 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
639 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
640 |
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
|
641 |
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
|
642 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
643 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
644 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
645 |
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
|
646 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
647 |
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
|
648 |
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
|
649 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
650 |
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
|
651 |
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
|
652 |
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
|
653 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
654 |
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
|
655 |
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
|
656 |
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
|
657 |
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
|
658 |
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
|
659 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
660 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
661 |
lemma independent_strict_incl: |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
662 |
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
|
663 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
664 |
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
|
665 |
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
|
666 |
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
|
667 |
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
|
668 |
ultimately show ?thesis |
68687 | 669 |
using independent_backwards(1)[OF assms] by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
670 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
671 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
672 |
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
|
673 |
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
|
674 |
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
|
675 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
676 |
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
|
677 |
hence "Span K Us \<subset> Span K Vs" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
678 |
using independent_strict_incl[OF assms(1)] by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
679 |
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
|
680 |
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
|
681 |
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
|
682 |
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
|
683 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
684 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
685 |
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
|
686 |
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
|
687 |
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
|
688 |
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
|
689 |
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
|
690 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
691 |
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
|
692 |
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
|
693 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
694 |
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
|
695 |
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
|
696 |
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
|
697 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
698 |
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
|
699 |
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
|
700 |
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
|
701 |
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
|
702 |
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
|
703 |
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
|
704 |
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
|
705 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
706 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
707 |
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
|
708 |
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
|
709 |
case Nil thus ?case |
68687 | 710 |
using Span_subgroup_props(2)[OF independent_in_carrier[of K Vs]] by simp |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
711 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
712 |
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
|
713 |
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
|
714 |
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
|
715 |
"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
|
716 |
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
|
717 |
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
|
718 |
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
|
719 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
720 |
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
|
721 |
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
|
722 |
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
|
723 |
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
|
724 |
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
|
725 |
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
|
726 |
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
|
727 |
and k: "k \<in> K" "(k \<otimes> u \<oplus> u') = v'" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
728 |
using line_extension_mem_iff[of _ _ u "Span K Us"] in_carrier(2-3)[THEN Span_subgroup_props(1)] |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
729 |
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
|
730 |
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
|
731 |
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
|
732 |
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
|
733 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
734 |
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
|
735 |
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
|
736 |
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
|
737 |
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
|
738 |
hence "k \<otimes> u \<in> Span K (Us @ Vs)" |
68687 | 739 |
using Span_subgroup_props(4)[OF in_carrier(2) u'(1)] v'(1) |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
740 |
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
|
741 |
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
|
742 |
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
|
743 |
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
|
744 |
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
|
745 |
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
|
746 |
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
|
747 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
748 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
749 |
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
|
750 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
751 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
752 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
753 |
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
|
754 |
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
|
755 |
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
|
756 |
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
|
757 |
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
|
758 |
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
|
759 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
760 |
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
|
761 |
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
|
762 |
"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
|
763 |
using Cons(2-3)[THEN independent_in_carrier] by auto |
68687 | 764 |
hence "Span K Us \<subseteq> Span K (u # Us)" |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
765 |
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
|
766 |
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
|
767 |
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
|
768 |
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
|
769 |
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
|
770 |
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
|
771 |
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
|
772 |
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
|
773 |
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
|
774 |
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
|
775 |
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
|
776 |
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
|
777 |
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
|
778 |
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
|
779 |
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
|
780 |
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
|
781 |
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
|
782 |
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
|
783 |
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
|
784 |
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
|
785 |
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
|
786 |
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
|
787 |
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
|
788 |
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
|
789 |
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
|
790 |
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
|
791 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
792 |
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
|
793 |
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
|
794 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
795 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
796 |
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
|
797 |
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
|
798 |
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
|
799 |
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
|
800 |
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
|
801 |
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
|
802 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
803 |
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
|
804 |
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
|
805 |
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
|
806 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
807 |
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
|
808 |
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
|
809 |
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
|
810 |
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
|
811 |
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
|
812 |
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
|
813 |
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
|
814 |
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
|
815 |
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
|
816 |
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
|
817 |
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
|
818 |
hence "set (take (length Us) Ks') \<subseteq> { \<zero> }" |
68687 | 819 |
using Cons(1)[OF Ks' _ independent_backwards(2)[OF Cons(4)]] by simp |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
820 |
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
|
821 |
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
|
822 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
823 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
824 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
825 |
lemma non_trivial_combine_imp_dependent: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
826 |
assumes "set Ks \<subseteq> K" and "combine Ks Us = \<zero>" and "\<not> set (take (length Us) Ks) \<subseteq> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
827 |
shows "dependent K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
828 |
using independent_imp_trivial_combine[OF _ assms(1-2)] assms(3) by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
829 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
830 |
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
|
831 |
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
|
832 |
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
|
833 |
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
|
834 |
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
|
835 |
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
|
836 |
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
|
837 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
838 |
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
|
839 |
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
|
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 |
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
|
842 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
843 |
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
|
844 |
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
|
845 |
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
|
846 |
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
|
847 |
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
|
848 |
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
|
849 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
850 |
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
|
851 |
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
|
852 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
853 |
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
|
854 |
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
|
855 |
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
|
856 |
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
|
857 |
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
|
858 |
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
|
859 |
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
|
860 |
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
|
861 |
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
|
862 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
863 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
864 |
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
|
865 |
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
|
866 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
867 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
868 |
corollary dependent_imp_non_trivial_combine: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
869 |
assumes "set Us \<subseteq> carrier R" and "dependent K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
870 |
obtains Ks where "length Ks = length Us" "combine Ks Us = \<zero>" "set Ks \<subseteq> K" "set Ks \<noteq> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
871 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
872 |
obtain Ks |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
873 |
where Ks: "set Ks \<subseteq> carrier R" "set Ks \<subseteq> K" "combine Ks Us = \<zero>" "\<not> set (take (length Us) Ks) \<subseteq> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
874 |
using trivial_combine_imp_independent[OF assms(1)] assms(2) subring_props(1) by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
875 |
obtain Ks' |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
876 |
where Ks': "set (take (length Us) Ks) \<subseteq> set Ks'" "set Ks' \<subseteq> set (take (length Us) Ks) \<union> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
877 |
"length Ks' = length Us" "combine Ks' Us = \<zero>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
878 |
using combine_normalize[OF Ks(1) assms(1) Ks(3)] by metis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
879 |
have "set (take (length Us) Ks) \<subseteq> set Ks" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
880 |
by (simp add: set_take_subset) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
881 |
hence "set Ks' \<subseteq> K" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
882 |
using Ks(2) Ks'(2) subring_props(2) Un_commute by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
883 |
moreover have "set Ks' \<noteq> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
884 |
using Ks'(1) Ks(4) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
885 |
ultimately show thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
886 |
using that Ks' by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
887 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
888 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
889 |
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
|
890 |
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
|
891 |
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
|
892 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
893 |
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
|
894 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
895 |
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
|
896 |
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
|
897 |
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
|
898 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
899 |
moreover |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
900 |
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
|
901 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
902 |
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
|
903 |
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
|
904 |
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
|
905 |
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
|
906 |
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
|
907 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
908 |
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
|
909 |
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
|
910 |
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
|
911 |
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
|
912 |
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
|
913 |
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
|
914 |
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
|
915 |
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
|
916 |
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
|
917 |
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
|
918 |
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
|
919 |
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
|
920 |
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
|
921 |
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
|
922 |
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
|
923 |
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
|
924 |
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
|
925 |
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
|
926 |
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
|
927 |
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
|
928 |
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
|
929 |
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
|
930 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
931 |
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
|
932 |
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
|
933 |
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
|
934 |
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
|
935 |
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
|
936 |
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
|
937 |
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
|
938 |
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
|
939 |
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
|
940 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
941 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
942 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
943 |
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
|
944 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
945 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
946 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
947 |
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
|
948 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
949 |
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
|
950 |
"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
|
951 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
952 |
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
|
953 |
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
|
954 |
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
|
955 |
"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
|
956 |
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
|
957 |
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
|
958 |
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
|
959 |
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
|
960 |
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
|
961 |
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
|
962 |
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
|
963 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
964 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
965 |
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
|
966 |
"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
|
967 |
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
|
968 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
969 |
(* |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
970 |
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
|
971 |
"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
|
972 |
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
|
973 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
974 |
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
|
975 |
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
|
976 |
*) |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
977 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
978 |
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
|
979 |
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
|
980 |
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
|
981 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
982 |
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
|
983 |
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
|
984 |
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
|
985 |
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
|
986 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
987 |
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
|
988 |
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
|
989 |
by (metis list.set_intros(1) split_list) |
68687 | 990 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
991 |
have in_carrier: "u \<in> carrier R" "set Us \<subseteq> carrier R" |
68687 | 992 |
using independent_in_carrier[OF Cons(2)] by auto |
993 |
||
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
994 |
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
|
995 |
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
|
996 |
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
|
997 |
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
|
998 |
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
|
999 |
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
|
1000 |
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
|
1001 |
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
|
1002 |
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
|
1003 |
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
|
1004 |
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
|
1005 |
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
|
1006 |
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
|
1007 |
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
|
1008 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1009 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1010 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1011 |
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
|
1012 |
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
|
1013 |
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
|
1014 |
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
|
1015 |
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
|
1016 |
proof (induct "length Us'" arbitrary: Us' Us) |
68687 | 1017 |
case 0 thus ?case by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1018 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1019 |
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
|
1020 |
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
|
1021 |
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
|
1022 |
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
|
1023 |
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
|
1024 |
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
|
1025 |
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
|
1026 |
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
|
1027 |
"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
|
1028 |
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
|
1029 |
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
|
1030 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1031 |
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
|
1032 |
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
|
1033 |
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
|
1034 |
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
|
1035 |
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
|
1036 |
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
|
1037 |
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
|
1038 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1039 |
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
|
1040 |
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
|
1041 |
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
|
1042 |
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
|
1043 |
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
|
1044 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1045 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1046 |
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
|
1047 |
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
|
1048 |
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
|
1049 |
proof - |
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 "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
|
1051 |
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
|
1052 |
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
|
1053 |
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
|
1054 |
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
|
1055 |
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
|
1056 |
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
|
1057 |
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
|
1058 |
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
|
1059 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1060 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1061 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1062 |
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
|
1063 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1064 |
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
|
1065 |
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
|
1066 |
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
|
1067 |
(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
|
1068 |
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
|
1069 |
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
|
1070 |
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
|
1071 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1072 |
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
|
1073 |
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
|
1074 |
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
|
1075 |
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
|
1076 |
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
|
1077 |
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
|
1078 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1079 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1080 |
lemma dimension_zero: "dimension 0 K E \<Longrightarrow> E = { \<zero> }" |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1081 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1082 |
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
|
1083 |
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
|
1084 |
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
|
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 |
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 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1089 |
lemma dimension_one [iff]: "dimension 1 K K" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1090 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1091 |
have "K = Span K [ \<one> ]" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1092 |
using line_extension_mem_iff[of _ K \<one> "{ \<zero> }"] subfieldE(3)[OF K] by (auto simp add: rev_subsetD) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1093 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1094 |
using dimension.Suc_dim[OF one_closed _ dimension.zero_dim, of K] subfieldE(6)[OF K] by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1095 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1096 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1097 |
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
|
1098 |
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
|
1099 |
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
|
1100 |
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
|
1101 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1102 |
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
|
1103 |
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
|
1104 |
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
|
1105 |
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
|
1106 |
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
|
1107 |
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
|
1108 |
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
|
1109 |
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
|
1110 |
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
|
1111 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1112 |
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
|
1113 |
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
|
1114 |
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
|
1115 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1116 |
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
|
1117 |
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
|
1118 |
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
|
1119 |
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
|
1120 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1121 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1122 |
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
|
1123 |
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
|
1124 |
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
|
1125 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1126 |
{ 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
|
1127 |
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
|
1128 |
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
|
1129 |
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
|
1130 |
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
|
1131 |
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
|
1132 |
} 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
|
1133 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1134 |
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
|
1135 |
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
|
1136 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1137 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1138 |
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
|
1139 |
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
|
1140 |
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
|
1141 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1142 |
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
|
1143 |
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
|
1144 |
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
|
1145 |
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
|
1146 |
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
|
1147 |
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
|
1148 |
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
|
1149 |
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
|
1150 |
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
|
1151 |
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
|
1152 |
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
|
1153 |
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
|
1154 |
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
|
1155 |
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
|
1156 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1157 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1158 |
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
|
1159 |
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
|
1160 |
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
|
1161 |
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
|
1162 |
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
|
1163 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1164 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1165 |
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
|
1166 |
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
|
1167 |
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
|
1168 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1169 |
{ 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
|
1170 |
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
|
1171 |
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
|
1172 |
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
|
1173 |
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
|
1174 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1175 |
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
|
1176 |
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
|
1177 |
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
|
1178 |
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
|
1179 |
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
|
1180 |
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
|
1181 |
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
|
1182 |
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
|
1183 |
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
|
1184 |
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
|
1185 |
where "length (Vs @ (v # Us)) = n" "independent K (Vs @ (v # Us))" "Span K (Vs @ (v # Us)) = E" |
68687 | 1186 |
using step(3)[of "v # Us"] step(1-2,4-6) v by auto |
68569
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 ?case |
68687 | 1188 |
by (metis append.assoc append_Cons append_Nil) |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1189 |
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
|
1190 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1191 |
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
|
1192 |
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
|
1193 |
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
|
1194 |
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
|
1195 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1196 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1197 |
lemma filter_base: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1198 |
assumes "set Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1199 |
obtains Vs where "set Vs \<subseteq> carrier R" and "independent K Vs" and "Span K Vs = Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1200 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1201 |
from \<open>set Us \<subseteq> carrier R\<close> have "\<exists>Vs. independent K Vs \<and> Span K Vs = Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1202 |
proof (induction Us) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1203 |
case Nil thus ?case by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1204 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1205 |
case (Cons u Us) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1206 |
then obtain Vs where Vs: "independent K Vs" "Span K Vs = Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1207 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1208 |
show ?case |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1209 |
proof (cases "u \<in> Span K Us") |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1210 |
case True |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1211 |
hence "Span K (u # Us) = Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1212 |
using Span_base_incl mono_Span_subset |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1213 |
by (metis Cons.prems insert_subset list.simps(15) subset_antisym) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1214 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1215 |
using Vs by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1216 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1217 |
case False |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1218 |
hence "Span K (u # Vs) = Span K (u # Us)" and "independent K (u # Vs)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1219 |
using li_Cons[of u K Vs] Cons(2) Vs by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1220 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1221 |
by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1222 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1223 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1224 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1225 |
using independent_in_carrier that by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1226 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1227 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1228 |
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
|
1229 |
"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
|
1230 |
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
|
1231 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1232 |
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
|
1233 |
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
|
1234 |
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
|
1235 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1236 |
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
|
1237 |
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
|
1238 |
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
|
1239 |
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
|
1240 |
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
|
1241 |
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
|
1242 |
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
|
1243 |
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
|
1244 |
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
|
1245 |
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
|
1246 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1247 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1248 |
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
|
1249 |
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
|
1250 |
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
|
1251 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1252 |
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
|
1253 |
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
|
1254 |
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
|
1255 |
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
|
1256 |
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
|
1257 |
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
|
1258 |
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
|
1259 |
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
|
1260 |
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
|
1261 |
hence "Span K Us \<inter> (Span K (Vs @ Bs)) \<subseteq> Span K Bs" |
68687 | 1262 |
using Bs(4) Us(3) Vs(3) mono_Span_append(1)[OF _ Bs(1), of Us] by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1263 |
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
|
1264 |
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
|
1265 |
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
|
1266 |
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
|
1267 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1268 |
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
|
1269 |
using independent_append[OF independent_split(2)[OF Us(2)] Vs(2)] Us(1) Vs(1) Bs(2) |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1270 |
dimension_independent[of K "Us @ (Vs @ Bs)"] by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1271 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1272 |
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
|
1273 |
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
|
1274 |
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
|
1275 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1276 |
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
|
1277 |
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
|
1278 |
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
|
1279 |
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
|
1280 |
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
|
1281 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1282 |
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
|
1283 |
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
|
1284 |
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
|
1285 |
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
|
1286 |
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
|
1287 |
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
|
1288 |
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
|
1289 |
qed |
68687 | 1290 |
ultimately have "Span K (Us @ (Vs @ Bs)) = E <+>\<^bsub>R\<^esub> F" |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1291 |
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
|
1292 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1293 |
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
|
1294 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1295 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1296 |
end (* of fixed K context. *) |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1297 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1298 |
end (* of ring context. *) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1299 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1300 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1301 |
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
|
1302 |
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
|
1303 |
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
|
1304 |
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
|
1305 |
proof - |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1306 |
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
|
1307 |
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
|
1308 |
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
|
1309 |
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
|
1310 |
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
|
1311 |
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
|
1312 |
using Us(1) u(1) by (induct Us) (auto) |
68687 | 1313 |
|
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1314 |
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
|
1315 |
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
|
1316 |
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
|
1317 |
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
|
1318 |
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
|
1319 |
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
|
1320 |
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
|
1321 |
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
|
1322 |
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
|
1323 |
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
|
1324 |
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
|
1325 |
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
|
1326 |
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
|
1327 |
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
|
1328 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1329 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1330 |
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
|
1331 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1332 |
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
|
1333 |
then obtain f where f: "f \<in> F" "v = f \<otimes> u \<oplus> \<zero>" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1334 |
using u(1,3) line_extension_mem_iff by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1335 |
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
|
1336 |
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
|
1337 |
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
|
1338 |
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
|
1339 |
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
|
1340 |
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
|
1341 |
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
|
1342 |
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
|
1343 |
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
|
1344 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1345 |
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
|
1346 |
proof |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1347 |
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
|
1348 |
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
|
1349 |
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
|
1350 |
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
|
1351 |
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
|
1352 |
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
|
1353 |
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
|
1354 |
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
|
1355 |
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
|
1356 |
thus "v \<in> E" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1357 |
using u(3) line_extension_mem_iff by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1358 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1359 |
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
|
1360 |
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
|
1361 |
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
|
1362 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1363 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1364 |
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
|
1365 |
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
|
1366 |
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
|
1367 |
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
|
1368 |
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
|
1369 |
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
|
1370 |
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
|
1371 |
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
|
1372 |
next |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1373 |
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
|
1374 |
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
|
1375 |
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
|
1376 |
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
|
1377 |
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
|
1378 |
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
|
1379 |
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
|
1380 |
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
|
1381 |
have "dimension n K (Span F [ v ])" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1382 |
using dimension_independent[OF li(1)] telescopic_base_aux[OF assms(1-3)] by simp |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1383 |
moreover have "dimension (n * m) K (Span F Vs')" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1384 |
using Suc(1) dimension_independent[OF li(2)] Vs(2) unfolding v by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1385 |
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
|
1386 |
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
|
1387 |
thus "dimension (n * Suc m) K E" |
68687 | 1388 |
using Span_append_eq_set_add[OF assms(2) li[THEN independent_in_carrier]] Vs(4) v by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1389 |
qed |
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1390 |
|
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1391 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1392 |
context ring_hom_ring |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1393 |
begin |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1394 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1395 |
lemma combine_hom: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1396 |
"\<lbrakk> set Ks \<subseteq> carrier R; set Us \<subseteq> carrier R \<rbrakk> \<Longrightarrow> combine (map h Ks) (map h Us) = h (R.combine Ks Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1397 |
by (induct Ks Us rule: R.combine.induct) (auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1398 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1399 |
lemma line_extension_hom: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1400 |
assumes "K \<subseteq> carrier R" "a \<in> carrier R" "E \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1401 |
shows "line_extension (h ` K) (h a) (h ` E) = h ` R.line_extension K a E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1402 |
using set_add_hom[OF homh R.r_coset_subset_G[OF assms(1-2)] assms(3)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1403 |
coset_hom(2)[OF ring_hom_in_hom(1)[OF homh] assms(1-2)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1404 |
unfolding R.line_extension_def S.line_extension_def |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1405 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1406 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1407 |
lemma Span_hom: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1408 |
assumes "K \<subseteq> carrier R" "set Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1409 |
shows "Span (h ` K) (map h Us) = h ` R.Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1410 |
using assms line_extension_hom R.Span_in_carrier by (induct Us) (auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1411 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1412 |
lemma inj_on_subgroup_iff_trivial_ker: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1413 |
assumes "subgroup H (add_monoid R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1414 |
shows "inj_on h H \<longleftrightarrow> a_kernel (R \<lparr> carrier := H \<rparr>) S h = { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1415 |
using group_hom.inj_on_subgroup_iff_trivial_ker[OF a_group_hom assms] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1416 |
unfolding a_kernel_def[of "R \<lparr> carrier := H \<rparr>" S h] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1417 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1418 |
corollary inj_on_Span_iff_trivial_ker: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1419 |
assumes "subfield K R" "set Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1420 |
shows "inj_on h (R.Span K Us) \<longleftrightarrow> a_kernel (R \<lparr> carrier := R.Span K Us \<rparr>) S h = { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1421 |
using inj_on_subgroup_iff_trivial_ker[OF R.Span_is_add_subgroup[OF assms]] . |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1422 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1423 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1424 |
context |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1425 |
fixes K :: "'a set" assumes K: "subfield K R" and one_zero: "\<one>\<^bsub>S\<^esub> \<noteq> \<zero>\<^bsub>S\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1426 |
begin |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1427 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1428 |
lemma inj_hom_preserves_independent: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1429 |
assumes "inj_on h (R.Span K Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1430 |
and "R.independent K Us" shows "independent (h ` K) (map h Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1431 |
proof (rule ccontr) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1432 |
have in_carrier: "set Us \<subseteq> carrier R" "set (map h Us) \<subseteq> carrier S" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1433 |
using R.independent_in_carrier[OF assms(2)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1434 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1435 |
assume ld: "dependent (h ` K) (map h Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1436 |
obtain Ks :: "'c list" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1437 |
where Ks: "length Ks = length Us" "combine Ks (map h Us) = \<zero>\<^bsub>S\<^esub>" "set Ks \<subseteq> h ` K" "set Ks \<noteq> { \<zero>\<^bsub>S\<^esub> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1438 |
using dependent_imp_non_trivial_combine[OF img_is_subfield(2)[OF K one_zero] in_carrier(2) ld] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1439 |
by (metis length_map) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1440 |
obtain Ks' where Ks': "set Ks' \<subseteq> K" "Ks = map h Ks'" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1441 |
using Ks(3) by (induct Ks) (auto, metis insert_subset list.simps(15,9)) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1442 |
hence "h (R.combine Ks' Us) = \<zero>\<^bsub>S\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1443 |
using combine_hom[OF _ in_carrier(1)] Ks(2) subfieldE(3)[OF K] by (metis subset_trans) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1444 |
moreover have "R.combine Ks' Us \<in> R.Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1445 |
using R.Span_eq_combine_set[OF K in_carrier(1)] Ks'(1) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1446 |
ultimately have "R.combine Ks' Us = \<zero>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1447 |
using assms hom_zero R.Span_subgroup_props(2)[OF K in_carrier(1)] by (auto simp add: inj_on_def) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1448 |
hence "set Ks' \<subseteq> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1449 |
using R.independent_imp_trivial_combine[OF K assms(2)] Ks' Ks(1) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1450 |
by (metis length_map order_refl take_all) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1451 |
hence "set Ks \<subseteq> { \<zero>\<^bsub>S\<^esub> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1452 |
unfolding Ks' using hom_zero by (induct Ks') (auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1453 |
hence "Ks = []" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1454 |
using Ks(4) by (metis set_empty2 subset_singletonD) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1455 |
hence "independent (h ` K) (map h Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1456 |
using independent.li_Nil Ks(1) by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1457 |
from \<open>dependent (h ` K) (map h Us)\<close> and this show False by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1458 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1459 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1460 |
corollary inj_hom_dimension: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1461 |
assumes "inj_on h E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1462 |
and "R.dimension n K E" shows "dimension n (h ` K) (h ` E)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1463 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1464 |
obtain Us |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1465 |
where Us: "set Us \<subseteq> carrier R" "R.independent K Us" "length Us = n" "R.Span K Us = E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1466 |
using R.exists_base[OF K assms(2)] by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1467 |
hence "dimension n (h ` K) (Span (h ` K) (map h Us))" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1468 |
using dimension_independent[OF inj_hom_preserves_independent[OF _ Us(2)]] assms(1) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1469 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1470 |
using Span_hom[OF subfieldE(3)[OF K] Us(1)] Us(4) by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1471 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1472 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1473 |
corollary rank_nullity_theorem: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1474 |
assumes "R.dimension n K E" and "R.dimension m K (a_kernel (R \<lparr> carrier := E \<rparr>) S h)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1475 |
shows "dimension (n - m) (h ` K) (h ` E)" |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1476 |
proof - |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1477 |
obtain Us |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1478 |
where Us: "set Us \<subseteq> carrier R" "R.independent K Us" "length Us = m" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1479 |
"R.Span K Us = a_kernel (R \<lparr> carrier := E \<rparr>) S h" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1480 |
using R.exists_base[OF K assms(2)] by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1481 |
obtain Vs |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1482 |
where Vs: "R.independent K (Vs @ Us)" "length (Vs @ Us) = n" "R.Span K (Vs @ Us) = E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1483 |
using R.complete_base[OF K assms(1) Us(2)] R.Span_base_incl[OF K Us(1)] Us(4) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1484 |
unfolding a_kernel_def' by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1485 |
have set_Vs: "set Vs \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1486 |
using R.independent_in_carrier[OF Vs(1)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1487 |
have "R.Span K Vs \<inter> a_kernel (R \<lparr> carrier := E \<rparr>) S h = { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1488 |
using R.independent_split[OF K Vs(1)] Us(4) by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1489 |
moreover have "R.Span K Vs \<subseteq> E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1490 |
using R.mono_Span_append(1)[OF K set_Vs Us(1)] Vs(3) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1491 |
ultimately have "a_kernel (R \<lparr> carrier := R.Span K Vs \<rparr>) S h \<subseteq> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1492 |
unfolding a_kernel_def' by (simp del: R.Span.simps, blast) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1493 |
hence "a_kernel (R \<lparr> carrier := R.Span K Vs \<rparr>) S h = { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1494 |
using R.Span_subgroup_props(2)[OF K set_Vs] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1495 |
unfolding a_kernel_def' by (auto simp del: R.Span.simps) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1496 |
hence "inj_on h (R.Span K Vs)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1497 |
using inj_on_Span_iff_trivial_ker[OF K set_Vs] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1498 |
moreover have "R.dimension (n - m) K (R.Span K Vs)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1499 |
using R.dimension_independent[OF R.independent_split(2)[OF K Vs(1)]] Vs(2) Us(3) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1500 |
ultimately have "dimension (n - m) (h ` K) (h ` (R.Span K Vs))" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1501 |
using assms(1) inj_hom_dimension by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1502 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1503 |
have "h ` E = h ` (R.Span K Vs <+>\<^bsub>R\<^esub> R.Span K Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1504 |
using R.Span_append_eq_set_add[OF K set_Vs Us(1)] Vs(3) by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1505 |
hence "h ` E = h ` (R.Span K Vs) <+>\<^bsub>S\<^esub> h ` (R.Span K Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1506 |
using R.Span_subgroup_props(1)[OF K] set_Vs Us(1) set_add_hom[OF homh] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1507 |
moreover have "h ` (R.Span K Us) = { \<zero>\<^bsub>S\<^esub> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1508 |
using R.space_subgroup_props(2)[OF K assms(1)] unfolding Us(4) a_kernel_def' by force |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1509 |
ultimately have "h ` E = h ` (R.Span K Vs) <+>\<^bsub>S\<^esub> { \<zero>\<^bsub>S\<^esub> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1510 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1511 |
hence "h ` E = h ` (R.Span K Vs)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1512 |
using R.Span_subgroup_props(1-2)[OF K set_Vs] unfolding set_add_def' by force |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1513 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1514 |
from \<open>dimension (n - m) (h ` K) (h ` (R.Span K Vs))\<close> and this show ?thesis by simp |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1515 |
qed |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1516 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1517 |
end (* of fixed K context. *) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1518 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1519 |
end (* of ring_hom_ring context. *) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1520 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1521 |
lemma (in ring_hom_ring) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1522 |
assumes "subfield K R" and "set Us \<subseteq> carrier R" and "\<one>\<^bsub>S\<^esub> \<noteq> \<zero>\<^bsub>S\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1523 |
and "independent (h ` K) (map h Us)" shows "R.independent K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1524 |
proof (rule ccontr) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1525 |
assume "R.dependent K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1526 |
then obtain Ks |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1527 |
where "length Ks = length Us" and "R.combine Ks Us = \<zero>" and "set Ks \<subseteq> K" and "set Ks \<noteq> { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1528 |
using R.dependent_imp_non_trivial_combine[OF assms(1-2)] by metis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1529 |
hence "combine (map h Ks) (map h Us) = \<zero>\<^bsub>S\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1530 |
using combine_hom[OF _ assms(2), of Ks] subfieldE(3)[OF assms(1)] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1531 |
moreover from \<open>set Ks \<subseteq> K\<close> have "set (map h Ks) \<subseteq> h ` K" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1532 |
by (induction Ks) (auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1533 |
moreover have "\<not> set (map h Ks) \<subseteq> { h \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1534 |
proof (rule ccontr) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1535 |
assume "\<not> \<not> set (map h Ks) \<subseteq> { h \<zero> }" then have "set (map h Ks) \<subseteq> { h \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1536 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1537 |
moreover from \<open>R.dependent K Us\<close> and \<open>length Ks = length Us\<close> have "Ks \<noteq> []" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1538 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1539 |
ultimately have "set (map h Ks) = { h \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1540 |
using subset_singletonD by fastforce |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1541 |
with \<open>set Ks \<subseteq> K\<close> have "set Ks = { \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1542 |
using inj_onD[OF _ _ _ subringE(2)[OF subfieldE(1)[OF assms(1)]], of h] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1543 |
img_is_subfield(1)[OF assms(1,3)] subset_singletonD |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1544 |
by (induction Ks) (auto simp add: subset_singletonD, fastforce) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1545 |
with \<open>set Ks \<noteq> { \<zero> }\<close> show False |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1546 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1547 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1548 |
with \<open>length Ks = length Us\<close> have "\<not> set (take (length (map h Us)) (map h Ks)) \<subseteq> { h \<zero> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1549 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1550 |
ultimately have "dependent (h ` K) (map h Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1551 |
using non_trivial_combine_imp_dependent[OF img_is_subfield(2)[OF assms(1,3)], of "map h Ks"] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1552 |
with \<open>independent (h ` K) (map h Us)\<close> show False |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1553 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1554 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1555 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1556 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1557 |
subsection \<open>Finite Dimension\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1558 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1559 |
definition (in ring) finite_dimension :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1560 |
where "finite_dimension K E \<longleftrightarrow> (\<exists>n. dimension n K E)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1561 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1562 |
abbreviation (in ring) infinite_dimension :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1563 |
where "infinite_dimension K E \<equiv> \<not> finite_dimension K E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1564 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1565 |
definition (in ring) dim :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1566 |
where "dim K E = (THE n. dimension n K E)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1567 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1568 |
locale subalgebra = subgroup V "add_monoid R" for K and V and R (structure) + |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1569 |
assumes smult_closed: "\<lbrakk> k \<in> K; v \<in> V \<rbrakk> \<Longrightarrow> k \<otimes> v \<in> V" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1570 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1571 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1572 |
subsubsection \<open>Basic Properties\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1573 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1574 |
lemma (in ring) unique_dimension: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1575 |
assumes "subfield K R" and "finite_dimension K E" shows "\<exists>!n. dimension n K E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1576 |
using assms(2) dimension_is_inj[OF assms(1)] unfolding finite_dimension_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1577 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1578 |
lemma (in ring) finite_dimensionI: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1579 |
assumes "dimension n K E" shows "finite_dimension K E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1580 |
using assms unfolding finite_dimension_def by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1581 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1582 |
lemma (in ring) finite_dimensionE: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1583 |
assumes "subfield K R" and "finite_dimension K E" shows "dimension ((dim over K) E) K E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1584 |
using theI'[OF unique_dimension[OF assms]] unfolding over_def dim_def by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1585 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1586 |
lemma (in ring) dimI: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1587 |
assumes "subfield K R" and "dimension n K E" shows "(dim over K) E = n" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1588 |
using finite_dimensionE[OF assms(1) finite_dimensionI] dimension_is_inj[OF assms(1)] assms(2) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1589 |
unfolding over_def dim_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1590 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1591 |
lemma (in ring) finite_dimensionE' [elim]: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1592 |
assumes "finite_dimension K E" and "\<And>n. dimension n K E \<Longrightarrow> P" shows P |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1593 |
using assms unfolding finite_dimension_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1594 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1595 |
lemma (in ring) Span_finite_dimension: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1596 |
assumes "subfield K R" and "set Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1597 |
shows "finite_dimension K (Span K Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1598 |
using filter_base[OF assms] finite_dimensionI[OF dimension_independent[of K]] by metis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1599 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1600 |
lemma (in ring) carrier_is_subalgebra: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1601 |
assumes "K \<subseteq> carrier R" shows "subalgebra K (carrier R) R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1602 |
using assms subalgebra.intro[OF add.group_incl_imp_subgroup[of "carrier R"], of K] add.group_axioms |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1603 |
unfolding subalgebra_axioms_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1604 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1605 |
lemma (in ring) subalgebra_in_carrier: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1606 |
assumes "subalgebra K V R" shows "V \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1607 |
using subgroup.subset[OF subalgebra.axioms(1)[OF assms]] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1608 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1609 |
lemma (in ring) subalgebra_inter: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1610 |
assumes "subalgebra K V R" and "subalgebra K V' R" shows "subalgebra K (V \<inter> V') R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1611 |
using add.subgroups_Inter_pair assms unfolding subalgebra_def subalgebra_axioms_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1612 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1613 |
lemma (in ring_hom_ring) img_is_subalgebra: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1614 |
assumes "K \<subseteq> carrier R" and "subalgebra K V R" shows "subalgebra (h ` K) (h ` V) S" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1615 |
proof (intro subalgebra.intro) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1616 |
have "group_hom (add_monoid R) (add_monoid S) h" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1617 |
using ring_hom_in_hom(2)[OF homh] R.add.group_axioms add.group_axioms |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1618 |
unfolding group_hom_def group_hom_axioms_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1619 |
thus "subgroup (h ` V) (add_monoid S)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1620 |
using group_hom.subgroup_img_is_subgroup[OF _ subalgebra.axioms(1)[OF assms(2)]] by force |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1621 |
next |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1622 |
show "subalgebra_axioms (h ` K) (h ` V) S" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1623 |
using R.subalgebra_in_carrier[OF assms(2)] subalgebra.axioms(2)[OF assms(2)] assms(1) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1624 |
unfolding subalgebra_axioms_def |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1625 |
by (auto, metis hom_mult image_eqI subset_iff) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1626 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1627 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1628 |
lemma (in ring) ideal_is_subalgebra: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1629 |
assumes "K \<subseteq> carrier R" "ideal I R" shows "subalgebra K I R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1630 |
using ideal.axioms(1)[OF assms(2)] ideal.I_l_closed[OF assms(2)] assms(1) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1631 |
unfolding subalgebra_def subalgebra_axioms_def additive_subgroup_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1632 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1633 |
lemma (in ring) Span_is_subalgebra: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1634 |
assumes "subfield K R" "set Us \<subseteq> carrier R" shows "subalgebra K (Span K Us) R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1635 |
using Span_smult_closed[OF assms] Span_is_add_subgroup[OF assms] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1636 |
unfolding subalgebra_def subalgebra_axioms_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1637 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1638 |
lemma (in ring) finite_dimension_imp_subalgebra: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1639 |
assumes "subfield K R" "finite_dimension K E" shows "subalgebra K E R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1640 |
using exists_base[OF assms(1) finite_dimensionE[OF assms]] Span_is_subalgebra[OF assms(1)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1641 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1642 |
lemma (in ring) subalgebra_Span_incl: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1643 |
assumes "subfield K R" and "subalgebra K V R" "set Us \<subseteq> V" shows "Span K Us \<subseteq> V" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1644 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1645 |
have "K <#> (set Us) \<subseteq> V" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1646 |
using subalgebra.smult_closed[OF assms(2)] assms(3) unfolding set_mult_def by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1647 |
moreover have "set Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1648 |
using subalgebra_in_carrier[OF assms(2)] assms(3) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1649 |
ultimately show ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1650 |
using subalgebra.axioms(1)[OF assms(2)] Span_min[OF assms(1)] by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1651 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1652 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1653 |
lemma (in ring) Span_subalgebra_minimal: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1654 |
assumes "subfield K R" "set Us \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1655 |
shows "Span K Us = \<Inter> { V. subalgebra K V R \<and> set Us \<subseteq> V }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1656 |
using Span_is_subalgebra[OF assms] Span_base_incl[OF assms] subalgebra_Span_incl[OF assms(1)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1657 |
by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1658 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1659 |
lemma (in ring) Span_subalgebraI: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1660 |
assumes "subfield K R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1661 |
and "subalgebra K E R" "set Us \<subseteq> E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1662 |
and "\<And>V. \<lbrakk> subalgebra K V R; set Us \<subseteq> V \<rbrakk> \<Longrightarrow> E \<subseteq> V" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1663 |
shows "E = Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1664 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1665 |
have "\<Inter> { V. subalgebra K V R \<and> set Us \<subseteq> V } = E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1666 |
using assms(2-4) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1667 |
thus "E = Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1668 |
using Span_subalgebra_minimal subalgebra_in_carrier[of K E] assms by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1669 |
qed |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1670 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1671 |
lemma (in ring) subalbegra_incl_imp_finite_dimension: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1672 |
assumes "subfield K R" and "finite_dimension K E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1673 |
and "subalgebra K V R" "V \<subseteq> E" shows "finite_dimension K V" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1674 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1675 |
obtain n where n: "dimension n K E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1676 |
using assms(2) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1677 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1678 |
define S where "S = { Us. set Us \<subseteq> V \<and> independent K Us }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1679 |
have "length ` S \<subseteq> {..n}" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1680 |
unfolding S_def using independent_length_le_dimension[OF assms(1) n] assms(4) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1681 |
moreover have "[] \<in> S" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1682 |
unfolding S_def by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1683 |
hence "length ` S \<noteq> {}" by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1684 |
ultimately obtain m where m: "m \<in> length ` S" and greatest: "\<And>k. k \<in> length ` S \<Longrightarrow> k \<le> m" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1685 |
by (meson Max_ge Max_in finite_atMost rev_finite_subset) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1686 |
then obtain Us where Us: "set Us \<subseteq> V" "independent K Us" "m = length Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1687 |
unfolding S_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1688 |
have "Span K Us = V" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1689 |
proof (rule ccontr) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1690 |
assume "\<not> Span K Us = V" then have "Span K Us \<subset> V" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1691 |
using subalgebra_Span_incl[OF assms(1,3) Us(1)] by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1692 |
then obtain v where v:"v \<in> V" "v \<notin> Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1693 |
by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1694 |
hence "independent K (v # Us)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1695 |
using independent.li_Cons[OF _ _ Us(2)] subalgebra_in_carrier[OF assms(3)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1696 |
hence "(v # Us) \<in> S" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1697 |
unfolding S_def using Us(1) v(1) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1698 |
hence "length (v # Us) \<le> m" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1699 |
using greatest by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1700 |
moreover have "length (v # Us) = Suc m" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1701 |
using Us(3) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1702 |
ultimately show False by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1703 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1704 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1705 |
using finite_dimensionI[OF dimension_independent[OF Us(2)]] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1706 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1707 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1708 |
lemma (in ring_hom_ring) infinite_dimension_hom: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1709 |
assumes "subfield K R" and "\<one>\<^bsub>S\<^esub> \<noteq> \<zero>\<^bsub>S\<^esub>" and "inj_on h E" and "subalgebra K E R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1710 |
shows "R.infinite_dimension K E \<Longrightarrow> infinite_dimension (h ` K) (h ` E)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1711 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1712 |
note subfield = img_is_subfield(2)[OF assms(1-2)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1713 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1714 |
assume "R.infinite_dimension K E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1715 |
show "infinite_dimension (h ` K) (h ` E)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1716 |
proof (rule ccontr) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1717 |
assume "\<not> infinite_dimension (h ` K) (h ` E)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1718 |
then obtain Vs where "set Vs \<subseteq> carrier S" and "Span (h ` K) Vs = h ` E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1719 |
using exists_base[OF subfield] by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1720 |
hence "set Vs \<subseteq> h ` E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1721 |
using Span_base_incl[OF subfield] by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1722 |
hence "\<exists>Us. set Us \<subseteq> E \<and> Vs = map h Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1723 |
by (induct Vs) (auto, metis insert_subset list.simps(9,15)) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1724 |
then obtain Us where "set Us \<subseteq> E" and "Vs = map h Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1725 |
by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1726 |
with \<open>Span (h ` K) Vs = h ` E\<close> have "h ` (R.Span K Us) = h ` E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1727 |
using R.subalgebra_in_carrier[OF assms(4)] Span_hom assms(1) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1728 |
moreover from \<open>set Us \<subseteq> E\<close> have "R.Span K Us \<subseteq> E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1729 |
using R.subalgebra_Span_incl assms(1-4) by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1730 |
ultimately have "R.Span K Us = E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1731 |
proof (auto simp del: R.Span.simps) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1732 |
fix a assume "a \<in> E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1733 |
with \<open>h ` (R.Span K Us) = h ` E\<close> obtain b where "b \<in> R.Span K Us" and "h a = h b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1734 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1735 |
with \<open>R.Span K Us \<subseteq> E\<close> and \<open>a \<in> E\<close> have "a = b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1736 |
using inj_onD[OF assms(3)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1737 |
with \<open>b \<in> R.Span K Us\<close> show "a \<in> R.Span K Us" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1738 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1739 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1740 |
with \<open>set Us \<subseteq> E\<close> have "R.finite_dimension K E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1741 |
using R.Span_finite_dimension[OF assms(1)] R.subalgebra_in_carrier[OF assms(4)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1742 |
with \<open>R.infinite_dimension K E\<close> show False |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1743 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1744 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1745 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1746 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1747 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1748 |
subsubsection \<open>Reformulation of some lemmas in this new language.\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1749 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1750 |
lemma (in ring) sum_space_dim: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1751 |
assumes "subfield K R" "finite_dimension K E" "finite_dimension K F" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1752 |
shows "finite_dimension K (E <+>\<^bsub>R\<^esub> F)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1753 |
and "((dim over K) (E <+>\<^bsub>R\<^esub> F)) = ((dim over K) E) + ((dim over K) F) - ((dim over K) (E \<inter> F))" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1754 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1755 |
obtain n m k where n: "dimension n K E" and m: "dimension m K F" and k: "dimension k K (E \<inter> F)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1756 |
using assms(2-3) subalbegra_incl_imp_finite_dimension[OF assms(1-2) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1757 |
subalgebra_inter[OF assms(2-3)[THEN finite_dimension_imp_subalgebra[OF assms(1)]]]] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1758 |
by (meson inf_le1 finite_dimension_def) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1759 |
hence "dimension (n + m - k) K (E <+>\<^bsub>R\<^esub> F)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1760 |
using dimension_sum_space[OF assms(1)] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1761 |
thus "finite_dimension K (E <+>\<^bsub>R\<^esub> F)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1762 |
and "((dim over K) (E <+>\<^bsub>R\<^esub> F)) = ((dim over K) E) + ((dim over K) F) - ((dim over K) (E \<inter> F))" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1763 |
using finite_dimensionI dimI[OF assms(1)] n m k by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1764 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1765 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1766 |
lemma (in ring) telescopic_base_dim: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1767 |
assumes "subfield K R" "subfield F R" and "finite_dimension K F" and "finite_dimension F E" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1768 |
shows "finite_dimension K E" and "(dim over K) E = ((dim over K) F) * ((dim over F) E)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1769 |
using telescopic_base[OF assms(1-2) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1770 |
finite_dimensionE[OF assms(1,3)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1771 |
finite_dimensionE[OF assms(2,4)]] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1772 |
dimI[OF assms(1)] finite_dimensionI |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
1773 |
by auto |
68569
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1774 |
|
68582 | 1775 |
end |