| author | paulson <lp15@cam.ac.uk> | 
| Tue, 25 Jan 2022 14:13:33 +0000 | |
| changeset 75008 | 43b3d5318d72 | 
| parent 73932 | fd21b4a93043 | 
| permissions | -rw-r--r-- | 
| 68582 | 1 | (* Title: HOL/Algebra/Subrings.thy | 
| 2 | Authors: Martin Baillon and 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 Subrings | 
| 
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 Ring RingHom QuotRing Multiplicative_Group | 
| 
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>Subrings\<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 | subsection \<open>Definitions\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | locale subring = | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | subgroup H "add_monoid R" + submonoid H R for H and R (structure) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | locale subcring = subring + | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | assumes sub_m_comm: "\<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = h2 \<otimes> h1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | locale subdomain = subcring + | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | assumes sub_one_not_zero [simp]: "\<one> \<noteq> \<zero>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | assumes subintegral: "\<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = \<zero> \<Longrightarrow> h1 = \<zero> \<or> h2 = \<zero>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | locale subfield = subdomain K R for K and R (structure) + | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 |   assumes subfield_Units: "Units (R \<lparr> carrier := K \<rparr>) = 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 | 25 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | subsection \<open>Basic Properties\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | subsubsection \<open>Subrings\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | lemma (in ring) subringI: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | assumes "H \<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 | 33 | and "\<one> \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | and "\<And>h. h \<in> H \<Longrightarrow> \<ominus> h \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<oplus> h2 \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | shows "subring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | using add.subgroupI[OF assms(1) _ assms(3, 5)] 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 | 39 | submonoid.intro[OF assms(1, 4, 2)] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | unfolding subring_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 | 41 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | lemma subringE: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | assumes "subring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | shows "H \<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 | 45 | and "\<zero>\<^bsub>R\<^esub> \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | and "\<one>\<^bsub>R\<^esub> \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 |     and "H \<noteq> {}"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | and "\<And>h. h \<in> H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | using subring.axioms[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 | 52 | unfolding submonoid_def subgroup_def a_inv_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 | 53 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | lemma (in ring) carrier_is_subring: "subring (carrier R) R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | by (simp add: subringI) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | lemma (in ring) subring_inter: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | assumes "subring I R" and "subring J R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | shows "subring (I \<inter> J) R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | using subringE[OF assms(1)] subringE[OF assms(2)] subringI[of "I \<inter> J"] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | |
| 
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 (in ring) subring_Inter: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 |   assumes "\<And>I. I \<in> S \<Longrightarrow> subring I R" and "S \<noteq> {}"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | shows "subring (\<Inter>S) R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | proof (rule subringI, auto simp add: assms subringE[of _ R]) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | fix x assume "\<forall>I \<in> S. x \<in> I" thus "x \<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 | 67 | using assms subringE(1)[of _ R] 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 | 68 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 70 | lemma (in ring) subring_is_ring: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 71 | assumes "subring H R" shows "ring (R \<lparr> carrier := H \<rparr>)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 72 | proof - | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 73 | interpret group "add_monoid (R \<lparr> carrier := H \<rparr>)" + monoid "R \<lparr> carrier := H \<rparr>" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 74 | using subgroup.subgroup_is_group[OF subring.axioms(1) add.is_group] assms | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 75 | submonoid.submonoid_is_monoid[OF subring.axioms(2) monoid_axioms] by auto | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 76 | show ?thesis | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 77 | using subringE(1)[OF assms] | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 78 | by (unfold_locales, simp_all add: subringE(1)[OF assms] add.m_comm subset_eq l_distr r_distr) | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 79 | qed | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | lemma (in ring) ring_incl_imp_subring: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | assumes "H \<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 | 83 | and "ring (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | shows "subring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | using group.group_incl_imp_subgroup[OF add.group_axioms, of H] assms(1) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | monoid.monoid_incl_imp_submonoid[OF monoid_axioms assms(1)] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | ring.axioms(1, 2)[OF assms(2)] abelian_group.a_group[of "R \<lparr> carrier := H \<rparr>"] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | unfolding subring_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 | 89 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | lemma (in ring) subring_iff: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | assumes "H \<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 | 92 | shows "subring H R \<longleftrightarrow> ring (R \<lparr> carrier := H \<rparr>)" | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 93 | using subring_is_ring ring_incl_imp_subring[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 | 94 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 95 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | subsubsection \<open>Subcrings\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | lemma (in ring) subcringI: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | assumes "subring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = h2 \<otimes> h1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | shows "subcring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | unfolding subcring_def subcring_axioms_def using assms by simp+ | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | lemma (in cring) subcringI': | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | assumes "subring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | shows "subcring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | using subcringI[OF assms] subringE(1)[OF assms] m_comm by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | |
| 
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 subcringE: | 
| 
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 "subcring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | shows "H \<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 "\<zero>\<^bsub>R\<^esub> \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | and "\<one>\<^bsub>R\<^esub> \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 |     and "H \<noteq> {}"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | and "\<And>h. h \<in> H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 = h2 \<otimes>\<^bsub>R\<^esub> h1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | using subringE[OF subcring.axioms(1)[OF assms]] subcring.sub_m_comm[OF assms] 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 | 120 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | lemma (in cring) carrier_is_subcring: "subcring (carrier R) 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 | by (simp add: subcringI' carrier_is_subring) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | lemma (in ring) subcring_inter: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | assumes "subcring I R" and "subcring J R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | shows "subcring (I \<inter> J) R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | using subcringE[OF assms(1)] subcringE[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 | 128 | subcringI[of "I \<inter> J"] subringI[of "I \<inter> J"] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | lemma (in ring) subcring_Inter: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 |   assumes "\<And>I. I \<in> S \<Longrightarrow> subcring I R" and "S \<noteq> {}"
 | 
| 
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 "subcring (\<Inter>S) R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | proof (rule subcringI) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | show "subring (\<Inter>S) R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | using subcring.axioms(1)[of _ R] subring_Inter[of S] 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 | 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 | fix h1 h2 assume h1: "h1 \<in> \<Inter>S" and h2: "h2 \<in> \<Inter>S" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | obtain S' where S': "S' \<in> S" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | using assms(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 | 140 | hence "h1 \<in> S'" "h2 \<in> S'" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | using h1 h2 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 | 142 | thus "h1 \<otimes> h2 = h2 \<otimes> h1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | using subcring.sub_m_comm[OF assms(1)[OF S']] 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 | 144 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | lemma (in ring) subcring_iff: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 | assumes "H \<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 | 148 | shows "subcring H R \<longleftrightarrow> cring (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | assume A: "subcring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | hence ring: "ring (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | using subring_iff[OF assms] subcring.axioms(1)[OF A] 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 | 153 | moreover have "comm_monoid (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | using monoid.monoid_comm_monoidI[OF ring.is_monoid[OF ring]] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | subcring.sub_m_comm[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 | 156 | ultimately show "cring (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | using cring_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 | 158 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 | assume A: "cring (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | hence "subring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 161 | using cring.axioms(1) subring_iff[OF assms] 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 | 162 | moreover have "comm_monoid (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | using A unfolding cring_def 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 | 164 | hence"\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = h2 \<otimes> h1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | using comm_monoid.m_comm[of "R \<lparr> carrier := H \<rparr>"] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | ultimately show "subcring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | unfolding subcring_def subcring_axioms_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 | 168 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | subsubsection \<open>Subdomains\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | lemma (in ring) subdomainI: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | assumes "subcring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | and "\<one> \<noteq> \<zero>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = \<zero> \<Longrightarrow> h1 = \<zero> \<or> h2 = \<zero>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | shows "subdomain H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | unfolding subdomain_def subdomain_axioms_def using assms 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 | 179 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | lemma (in domain) subdomainI': | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | assumes "subring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | shows "subdomain H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | proof (rule subdomainI[OF subcringI[OF assms]], 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 | 184 | show "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes> h2 = h2 \<otimes> h1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | using m_comm subringE(1)[OF assms] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | show "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H; h1 \<otimes> h2 = \<zero> \<rbrakk> \<Longrightarrow> (h1 = \<zero>) \<or> (h2 = \<zero>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | using integral subringE(1)[OF assms] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | lemma subdomainE: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | assumes "subdomain H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | shows "H \<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 | 193 | and "\<zero>\<^bsub>R\<^esub> \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | and "\<one>\<^bsub>R\<^esub> \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 |     and "H \<noteq> {}"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | and "\<And>h. h \<in> H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 = h2 \<otimes>\<^bsub>R\<^esub> h1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | and "\<And>h1 h2. \<lbrakk> h1 \<in> H; h2 \<in> H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 = \<zero>\<^bsub>R\<^esub> \<Longrightarrow> h1 = \<zero>\<^bsub>R\<^esub> \<or> h2 = \<zero>\<^bsub>R\<^esub>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | and "\<one>\<^bsub>R\<^esub> \<noteq> \<zero>\<^bsub>R\<^esub>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | using subcringE[OF subdomain.axioms(1)[OF assms]] assms | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | unfolding subdomain_def subdomain_axioms_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 | 204 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | lemma (in ring) subdomain_iff: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | assumes "H \<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 | 207 | shows "subdomain H R \<longleftrightarrow> domain (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | assume A: "subdomain H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | hence cring: "cring (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | using subcring_iff[OF assms] subdomain.axioms(1)[OF A] 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 | 212 | thus "domain (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | using domain.intro[OF cring] subdomain.subintegral[OF A] subdomain.sub_one_not_zero[OF A] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | unfolding domain_axioms_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 | 215 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | assume A: "domain (R \<lparr> carrier := H \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | hence subcring: "subcring H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | using subcring_iff[OF assms] unfolding domain_def 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 | 219 | thus "subdomain H R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | using subdomain.intro[OF subcring] domain.integral[OF A] domain.one_not_zero[OF A] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | unfolding subdomain_axioms_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 | 222 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | |
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 224 | lemma (in domain) subring_is_domain: | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 225 | assumes "subring H R" shows "domain (R \<lparr> carrier := H \<rparr>)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 226 | using subdomainI'[OF assms] unfolding subdomain_iff[OF subringE(1)[OF assms]] . | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 227 | |
| 70160 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 228 | (* NEW ====================== *) | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 229 | lemma (in ring) subdomain_is_domain: | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 230 | assumes "subdomain H R" shows "domain (R \<lparr> carrier := H \<rparr>)" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 231 | using assms unfolding subdomain_iff[OF subdomainE(1)[OF assms]] . | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 232 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | subsubsection \<open>Subfields\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | lemma (in ring) subfieldI: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 |   assumes "subcring K R" and "Units (R \<lparr> carrier := K \<rparr>) = 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 | 238 | shows "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 | 239 | proof (rule subfield.intro) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | show "subfield_axioms 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 | 241 | using assms(2) unfolding subfield_axioms_def . | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | show "subdomain 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 | 243 | proof (rule subdomainI[OF assms(1)], auto) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | have subM: "submonoid 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 | 245 | using subring.axioms(2)[OF subcring.axioms(1)[OF assms(1)]] . | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | show contr: "\<one> = \<zero> \<Longrightarrow> False" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | proof - | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | assume one_eq_zero: "\<one> = \<zero>" | 
| 
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 "\<one> \<in> K" and "\<one> \<otimes> \<one> = \<one>" | 
| 
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 submonoid.one_closed[OF subM] 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 | 252 | hence "\<one> \<in> Units (R \<lparr> carrier := K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | unfolding Units_def by (simp, blast) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | hence "\<one> \<noteq> \<zero>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | 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 | 256 | 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 | 257 | using one_eq_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 | 258 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | fix k1 k2 assume k1: "k1 \<in> K" and k2: "k2 \<in> K" "k2 \<noteq> \<zero>" and k12: "k1 \<otimes> k2 = \<zero>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | obtain k2' where k2': "k2' \<in> K" "k2' \<otimes> k2 = \<one>" "k2 \<otimes> k2' = \<one>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | using assms(2) k2 unfolding Units_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 | 263 | have "\<zero> = (k1 \<otimes> k2) \<otimes> k2'" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | using k12 k2'(1) submonoid.mem_carrier[OF subM] 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 | 265 | also have "... = k1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | using k1 k2(1) k2'(1,3) submonoid.mem_carrier[OF subM] by (simp 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 | 267 | finally have "\<zero> = k1" . | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | thus "k1 = \<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 | 269 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | lemma (in field) subfieldI': | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 |   assumes "subring K R" and "\<And>k. k \<in> K - { \<zero> } \<Longrightarrow> inv 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 | 274 | shows "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 | 275 | proof (rule subfieldI) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | show "subcring 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 | 277 | using subcringI[OF assms(1)] m_comm subringE(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 | 278 |   show "Units (R \<lparr> carrier := K \<rparr>) = 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 | 279 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 |     show "K - { \<zero> } \<subseteq> Units (R \<lparr> carrier := K \<rparr>)"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 |       fix k assume k: "k \<in> 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 | 283 | hence inv_k: "inv 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 | 284 | 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 | 285 |       moreover have "k \<in> carrier R - { \<zero> }" 
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | using subringE(1)[OF assms(1)] 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 | 287 | ultimately have "k \<otimes> inv k = \<one>" "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 | 288 | by (simp add: field_Units)+ | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | thus "k \<in> Units (R \<lparr> carrier := K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | unfolding Units_def using k inv_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 | 291 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 |     show "Units (R \<lparr> carrier := K \<rparr>) \<subseteq> 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 | 294 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | fix k assume k: "k \<in> Units (R \<lparr> carrier := K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | then obtain k' where k': "k' \<in> K" "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 | 297 | unfolding Units_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 | 298 | hence "k \<in> carrier R" and "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 | 299 | using k subringE(1)[OF assms(1)] unfolding Units_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 | 300 | hence "\<zero> = \<one>" 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 | 301 | using that k'(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 | 302 |       thus "k \<in> 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 | 303 | using k unfolding Units_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 | 304 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | lemma (in field) carrier_is_subfield: "subfield (carrier R) R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | by (auto intro: subfieldI[OF carrier_is_subcring] simp add: field_Units) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | lemma subfieldE: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | assumes "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 | 313 | shows "subring K R" and "subcring 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 | 314 | and "K \<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 | 315 | and "\<And>k1 k2. \<lbrakk> k1 \<in> K; k2 \<in> K \<rbrakk> \<Longrightarrow> k1 \<otimes>\<^bsub>R\<^esub> k2 = k2 \<otimes>\<^bsub>R\<^esub> k1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | and "\<And>k1 k2. \<lbrakk> k1 \<in> K; k2 \<in> K \<rbrakk> \<Longrightarrow> k1 \<otimes>\<^bsub>R\<^esub> k2 = \<zero>\<^bsub>R\<^esub> \<Longrightarrow> k1 = \<zero>\<^bsub>R\<^esub> \<or> k2 = \<zero>\<^bsub>R\<^esub>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | and "\<one>\<^bsub>R\<^esub> \<noteq> \<zero>\<^bsub>R\<^esub>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | using subdomain.axioms(1)[OF subfield.axioms(1)[OF assms]] subcring_def | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | subdomainE(1, 8, 9, 10)[OF subfield.axioms(1)[OF assms]] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | lemma (in ring) subfield_m_inv: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 |   assumes "subfield K R" and "k \<in> 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 | 323 |   shows "inv k \<in> K - { \<zero> }" and "k \<otimes> inv k = \<one>" and "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 | 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 | have K: "subring K R" "submonoid 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 | 326 | using subfieldE(1)[OF assms(1)] subring.axioms(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 | 327 | have monoid: "monoid (R \<lparr> carrier := K \<rparr>)" | 
| 
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 submonoid.submonoid_is_monoid[OF subring.axioms(2)[OF K(1)] is_monoid] . | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | have "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 | 331 | by (simp add: monoid_axioms) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | hence k: "k \<in> Units (R \<lparr> carrier := K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | using subfield.subfield_Units[OF assms(1)] assms(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 | 335 | hence unit_of_R: "k \<in> Units R" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | using assms(2) subringE(1)[OF subfieldE(1)[OF assms(1)]] unfolding Units_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 | 337 | have "inv\<^bsub>(R \<lparr> carrier := K \<rparr>)\<^esub> k \<in> Units (R \<lparr> carrier := K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | by (simp add: k monoid monoid.Units_inv_Units) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 |   hence "inv\<^bsub>(R \<lparr> carrier := K \<rparr>)\<^esub> k \<in> 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 | 340 | using subfield.subfield_Units[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 | 341 |   thus "inv k \<in> K - { \<zero> }" and "k \<otimes> inv k = \<one>" and "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 | 342 | using Units_l_inv[OF unit_of_R] Units_r_inv[OF unit_of_R] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | using monoid.m_inv_monoid_consistent[OF monoid_axioms k K(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 | 344 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | |
| 70160 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 346 | lemma (in ring) subfield_m_inv_simprule: | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 347 | assumes "subfield K R" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 348 |   shows "\<lbrakk> k \<in> K - { \<zero> }; a \<in> carrier R \<rbrakk> \<Longrightarrow> k \<otimes> a \<in> K \<Longrightarrow> a \<in> K"
 | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 349 | proof - | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 350 | note subring_props = subringE[OF subfieldE(1)[OF assms]] | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 351 | |
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 352 |   assume A: "k \<in> K - { \<zero> }" "a \<in> carrier R" "k \<otimes> a \<in> K"
 | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 353 | then obtain k' where k': "k' \<in> K" "k \<otimes> a = k'" by blast | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 354 | have inv_k: "inv k \<in> K" "inv k \<otimes> k = \<one>" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 355 | using subfield_m_inv[OF assms A(1)] by auto | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 356 | hence "inv k \<otimes> (k \<otimes> a) \<in> K" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 357 | using k' A(3) subring_props(6) by auto | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 358 | thus "a \<in> K" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 359 | using m_assoc[of "inv k" k a] A(2) inv_k subring_props(1) | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
70160diff
changeset | 360 | by (metis (no_types, opaque_lifting) A(1) Diff_iff l_one subsetCE) | 
| 70160 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 361 | qed | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 362 | |
| 68569 
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 (in ring) subfield_iff: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | shows "\<lbrakk> field (R \<lparr> carrier := K \<rparr>); K \<subseteq> carrier R \<rbrakk> \<Longrightarrow> 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 | 365 | and "subfield K R \<Longrightarrow> field (R \<lparr> carrier := K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | proof- | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | assume A: "field (R \<lparr> carrier := K \<rparr>)" "K \<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 | 368 | have "\<And>k1 k2. \<lbrakk> k1 \<in> K; k2 \<in> K \<rbrakk> \<Longrightarrow> k1 \<otimes> k2 = k2 \<otimes> k1" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | using comm_monoid.m_comm[OF cring.axioms(2)[OF fieldE(1)[OF A(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 | 370 | moreover have "subring 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 | 371 | using ring_incl_imp_subring[OF A(2) cring.axioms(1)[OF fieldE(1)[OF A(1)]]] . | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | ultimately have "subcring 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 | 373 | using subcringI 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 | 374 | thus "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 | 375 | using field.field_Units[OF A(1)] subfieldI by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | assume A: "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 | 378 | have cring: "cring (R \<lparr> carrier := K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | using subcring_iff[OF subringE(1)[OF subfieldE(1)[OF A]]] subfieldE(2)[OF A] 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 | 380 | thus "field (R \<lparr> carrier := K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | using cring.cring_fieldI[OF cring] subfield.subfield_Units[OF A] 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 | 382 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | lemma (in field) subgroup_mult_of : | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | assumes "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 | 386 |   shows "subgroup (K - {\<zero>}) (mult_of R)"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | proof (intro group.group_incl_imp_subgroup[OF field_mult_group]) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 |   show "K - {\<zero>} \<subseteq> carrier (mult_of R)"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | by (simp add: Diff_mono assms carrier_mult_of subfieldE(3)) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 |   show "group ((mult_of R) \<lparr> carrier := K - {\<zero>} \<rparr>)"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | using field.field_mult_group[OF subfield_iff(2)[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 | 392 | unfolding mult_of_def 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 | 393 | qed | 
| 
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 | subsection \<open>Subring Homomorphisms\<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 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | lemma (in ring) hom_imp_img_subring: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | assumes "h \<in> ring_hom R S" and "subring 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 | 400 | shows "ring (S \<lparr> carrier := h ` K, one := h \<one>, zero := h \<zero> \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68664diff
changeset | 402 | have [simp]: "h \<one> = \<one>\<^bsub>S\<^esub>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68664diff
changeset | 403 | using assms ring_hom_one by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | have "ring (R \<lparr> carrier := K \<rparr>)" | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 405 | by (simp add: assms(2) subring_is_ring) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | moreover have "h \<in> ring_hom (R \<lparr> carrier := K \<rparr>) S" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | using assms subringE(1)[OF assms (2)] unfolding ring_hom_def | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | apply simp | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | apply blast | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | done | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | 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 | 412 | using ring.ring_hom_imp_img_ring[of "R \<lparr> carrier := K \<rparr>" h S] 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 | 413 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | lemma (in ring_hom_ring) img_is_subring: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | assumes "subring K R" shows "subring (h ` K) S" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | proof - | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | have "ring (S \<lparr> carrier := h ` K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | using R.hom_imp_img_subring[OF homh assms] hom_zero hom_one 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 | 420 | moreover have "h ` K \<subseteq> carrier S" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 | using ring_hom_memE(1)[OF homh] subringE(1)[OF assms] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | 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 | 423 | using ring_incl_imp_subring 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 | 424 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | lemma (in ring_hom_ring) img_is_subfield: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | assumes "subfield K R" and "\<one>\<^bsub>S\<^esub> \<noteq> \<zero>\<^bsub>S\<^esub>" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | shows "inj_on h K" and "subfield (h ` K) S" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | proof - | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | have K: "K \<subseteq> carrier R" "subring K R" "subring (h ` K) S" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | using subfieldE(1)[OF assms(1)] subringE(1) img_is_subring by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | have field: "field (R \<lparr> carrier := K \<rparr>)" | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 433 | using R.subfield_iff(2) \<open>subfield K R\<close> by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 434 | moreover have ring: "ring (R \<lparr> carrier := K \<rparr>)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 435 | using K R.ring_axioms R.subring_is_ring by blast | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 436 | moreover have ringS: "ring (S \<lparr> carrier := h ` K \<rparr>)" | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 437 | using subring_is_ring K by simp | 
| 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 438 | ultimately have h: "h \<in> ring_hom (R \<lparr> carrier := K \<rparr>) (S \<lparr> carrier := h ` K \<rparr>)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | unfolding ring_hom_def apply auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 440 | using ring_hom_memE[OF homh] K | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | by (meson contra_subsetD)+ | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | hence ring_hom: "ring_hom_ring (R \<lparr> carrier := K \<rparr>) (S \<lparr> carrier := h ` K \<rparr>) h" | 
| 68664 
bd0df72c16d5
updated material concerning Algebra
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 443 | using ring_axioms ring ringS ring_hom_ringI2 by blast | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 |   have "h ` K \<noteq> { \<zero>\<^bsub>S\<^esub> }"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 | using subfieldE(1, 5)[OF assms(1)] subringE(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 | 446 | by (metis hom_one image_eqI singletonD) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 447 | thus "inj_on h K" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 448 | using ring_hom_ring.non_trivial_field_hom_imp_inj[OF ring_hom field] by auto | 
| 
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 | hence "h \<in> ring_iso (R \<lparr> carrier := K \<rparr>) (S \<lparr> carrier := h ` K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 451 | using h unfolding ring_iso_def bij_betw_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 | 452 | hence "field (S \<lparr> carrier := h ` K \<rparr>)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 453 | using field.ring_iso_imp_img_field[OF field, of h "S \<lparr> carrier := h ` K \<rparr>"] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 454 | thus "subfield (h ` K) S" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 455 | using S.subfield_iff[of "h ` K"] K(1) ring_hom_memE(1)[OF homh] 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 | 456 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | |
| 70160 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 458 | (* NEW ========================================================================== *) | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 459 | lemma (in ring_hom_ring) induced_ring_hom: | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 460 | assumes "subring K R" shows "ring_hom_ring (R \<lparr> carrier := K \<rparr>) S h" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 461 | proof - | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 462 | have "h \<in> ring_hom (R \<lparr> carrier := K \<rparr>) S" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 463 | using homh subringE(1)[OF assms] unfolding ring_hom_def | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 464 | by (auto, meson hom_mult hom_add subsetCE)+ | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 465 | thus ?thesis | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 466 | using R.subring_is_ring[OF assms] ring_axioms | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 467 | unfolding ring_hom_ring_def ring_hom_ring_axioms_def by auto | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 468 | qed | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 469 | |
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 470 | (* NEW ========================================================================== *) | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 471 | lemma (in ring_hom_ring) inj_on_subgroup_iff_trivial_ker: | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 472 | assumes "subring K R" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 473 |   shows "inj_on h K \<longleftrightarrow> a_kernel (R \<lparr> carrier := K \<rparr>) S h = { \<zero> }"
 | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 474 | using ring_hom_ring.inj_iff_trivial_ker[OF induced_ring_hom[OF assms]] by simp | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 475 | |
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 476 | lemma (in ring_hom_ring) inv_ring_hom: | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 477 | assumes "inj_on h K" and "subring K R" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 478 | shows "ring_hom_ring (S \<lparr> carrier := h ` K \<rparr>) R (inv_into K h)" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 479 | proof (intro ring_hom_ringI[OF _ R.ring_axioms], auto) | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 480 | show "ring (S \<lparr> carrier := h ` K \<rparr>)" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 481 | using subring_is_ring[OF img_is_subring[OF assms(2)]] . | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 482 | next | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 483 | show "inv_into K h \<one>\<^bsub>S\<^esub> = \<one>\<^bsub>R\<^esub>" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 484 | using assms(1) subringE(3)[OF assms(2)] hom_one by (simp add: inv_into_f_eq) | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 485 | next | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 486 | fix k1 k2 | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 487 | assume k1: "k1 \<in> K" and k2: "k2 \<in> K" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 488 | with \<open>k1 \<in> K\<close> show "inv_into K h (h k1) \<in> carrier R" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 489 | using assms(1) subringE(1)[OF assms(2)] by (simp add: subset_iff) | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 490 | |
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 491 | from \<open>k1 \<in> K\<close> and \<open>k2 \<in> K\<close> | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 492 | have "h k1 \<oplus>\<^bsub>S\<^esub> h k2 = h (k1 \<oplus>\<^bsub>R\<^esub> k2)" and "k1 \<oplus>\<^bsub>R\<^esub> k2 \<in> K" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 493 | and "h k1 \<otimes>\<^bsub>S\<^esub> h k2 = h (k1 \<otimes>\<^bsub>R\<^esub> k2)" and "k1 \<otimes>\<^bsub>R\<^esub> k2 \<in> K" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 494 | using subringE(1,6,7)[OF assms(2)] by (simp add: subset_iff)+ | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 495 | thus "inv_into K h (h k1 \<oplus>\<^bsub>S\<^esub> h k2) = inv_into K h (h k1) \<oplus>\<^bsub>R\<^esub> inv_into K h (h k2)" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 496 | and "inv_into K h (h k1 \<otimes>\<^bsub>S\<^esub> h k2) = inv_into K h (h k1) \<otimes>\<^bsub>R\<^esub> inv_into K h (h k2)" | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 497 | using assms(1) k1 k2 by simp+ | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 498 | qed | 
| 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 paulson <lp15@cam.ac.uk> parents: 
69122diff
changeset | 499 | |
| 68582 | 500 | end |