| author | wenzelm | 
| Wed, 09 Dec 2020 15:53:45 +0100 | |
| changeset 72857 | a9e091ccd450 | 
| parent 68582 | b9b9e2985878 | 
| child 81438 | 95c9af7483b1 | 
| permissions | -rw-r--r-- | 
| 68582 | 1 | (* Title: HOL/Algebra/Generated_Fields.thy | 
| 2 | Author: Martin Baillon | |
| 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 Generated_Fields | 
| 
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 Generated_Rings Subrings 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 | inductive_set | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 |   generate_field :: "('a, 'b) ring_scheme \<Rightarrow> 'a set \<Rightarrow> 'a set"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | for R and H where | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | one : "\<one>\<^bsub>R\<^esub> \<in> generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | | incl : "h \<in> H \<Longrightarrow> h \<in> generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | | a_inv: "h \<in> generate_field R H \<Longrightarrow> \<ominus>\<^bsub>R\<^esub> h \<in> generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | | m_inv: "\<lbrakk> h \<in> generate_field R H; h \<noteq> \<zero>\<^bsub>R\<^esub> \<rbrakk> \<Longrightarrow> inv\<^bsub>R\<^esub> h \<in> generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | | eng_add : "\<lbrakk> h1 \<in> generate_field R H; h2 \<in> generate_field R H \<rbrakk> \<Longrightarrow> h1 \<oplus>\<^bsub>R\<^esub> h2 \<in> generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | | eng_mult: "\<lbrakk> h1 \<in> generate_field R H; h2 \<in> generate_field R H \<rbrakk> \<Longrightarrow> h1 \<otimes>\<^bsub>R\<^esub> h2 \<in> generate_field R H" | 
| 
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 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | subsection\<open>Basic Properties of Generated Rings - First Part\<close> | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | lemma (in field) generate_field_in_carrier: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | 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 | 24 | shows "h \<in> generate_field R H \<Longrightarrow> h \<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 | 25 | apply (induction rule: generate_field.induct) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | using assms 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 | 27 | 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 | 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 | lemma (in field) generate_field_incl: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | 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 | 31 | shows "generate_field R 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 | 32 | using generate_field_in_carrier[OF assms] by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | lemma (in field) zero_in_generate: "\<zero>\<^bsub>R\<^esub> \<in> generate_field R 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 | using one a_inv generate_field.eng_add one_closed r_neg | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | by metis | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | lemma (in field) generate_field_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 | 39 | 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 | 40 | shows "subfield (generate_field R 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 | 41 | proof (intro subfieldI', simp_all add: 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 | 42 | show "subring (generate_field R 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 | 43 | by (auto intro: subringI[of "generate_field R H"] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | simp add: eng_add a_inv eng_mult one generate_field_in_carrier[OF assms]) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | lemma (in field) generate_field_is_add_subgroup: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | 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 | 49 | shows "subgroup (generate_field R H) (add_monoid R)" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | using subring.axioms(1)[OF subfieldE(1)[OF generate_field_is_subfield[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 | 51 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | lemma (in field) generate_field_is_field : | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | 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 | 54 | shows "field (R \<lparr> carrier := generate_field R 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 | 55 | using subfield_iff generate_field_is_subfield 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 | 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 field) generate_field_min_subfield1: | 
| 
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 "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 | 59 | and "subfield E R" "H \<subseteq> E" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | shows "generate_field R H \<subseteq> E" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | fix h | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | assume h: "h \<in> generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | show "h \<in> E" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | using h and assms(3) and subfield_m_inv[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 | 66 | by (induct rule: generate_field.induct) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | (auto simp add: subringE(3,5-7)[OF subfieldE(1)[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 | 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 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | lemma (in field) generate_fieldI: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | 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 | 72 | and "subfield E R" "H \<subseteq> E" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | and "\<And>K. \<lbrakk> subfield K R; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | shows "E = generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | show "E \<subseteq> generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | using assms generate_field_is_subfield generate_field.incl by (metis subset_iff) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | show "generate_field R H \<subseteq> E" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | using generate_field_min_subfield1[OF assms(1-3)] by simp | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | lemma (in field) generate_fieldE: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | assumes "H \<subseteq> carrier R" and "E = generate_field R H" | 
| 
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 "subfield E R" and "H \<subseteq> E" and "\<And>K. \<lbrakk> subfield K R; H \<subseteq> K \<rbrakk> \<Longrightarrow> E \<subseteq> K" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | proof - | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | show "subfield E R" using assms generate_field_is_subfield 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 | 87 | show "H \<subseteq> E" using assms(2) by (simp add: generate_field.incl subsetI) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | show "\<And>K. subfield K R \<Longrightarrow> H \<subseteq> K \<Longrightarrow> E \<subseteq> K" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | using assms generate_field_min_subfield1 by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | lemma (in field) generate_field_min_subfield2: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | 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 | 94 |   shows "generate_field R H = \<Inter>{K. subfield K R \<and> H \<subseteq> K}"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | have "subfield (generate_field R H) R \<and> H \<subseteq> generate_field R H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | by (simp add: assms generate_fieldE(2) generate_field_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 | 98 |   thus "\<Inter>{K. subfield K R \<and> H \<subseteq> K} \<subseteq> generate_field R H" 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 | 99 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | have "\<And>K. subfield K R \<and> H \<subseteq> K \<Longrightarrow> generate_field R H \<subseteq> K" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | by (simp add: assms generate_field_min_subfield1) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 |   thus "generate_field R H \<subseteq> \<Inter>{K. subfield K R \<and> H \<subseteq> K}" 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 | 103 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | lemma (in field) mono_generate_field: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | assumes "I \<subseteq> J" and "J \<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 | 107 | shows "generate_field R I \<subseteq> generate_field R J" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | proof- | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | have "I \<subseteq> generate_field R J " | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | using assms generate_fieldE(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 | 111 | thus "generate_field R I \<subseteq> generate_field R J" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | using generate_field_min_subfield1[of I "generate_field R J"] assms generate_field_is_subfield[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 | 113 | 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 | 114 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | lemma (in field) subfield_gen_incl : | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | assumes "subfield 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 | 119 | and "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 | 120 | and "I \<subseteq> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | and "I \<subseteq> K" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | shows "generate_field (R\<lparr>carrier := K\<rparr>) I \<subseteq> generate_field (R\<lparr>carrier := H\<rparr>) I" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | proof | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 |   {fix J assume J_def : "subfield J R" "I \<subseteq> J"
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | have "generate_field (R \<lparr> carrier := J \<rparr>) I \<subseteq> J" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | using field.mono_generate_field[of "(R\<lparr>carrier := J\<rparr>)" I J] subfield_iff(2)[OF J_def(1)] | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | field.generate_field_in_carrier[of "R\<lparr>carrier := J\<rparr>"] field_axioms J_def | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | by 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 | note incl_HK = this | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 |   {fix x have "x \<in> generate_field (R\<lparr>carrier := K\<rparr>) I \<Longrightarrow> x \<in> generate_field (R\<lparr>carrier := H\<rparr>) I" 
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | proof (induction rule : generate_field.induct) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | case one | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | have "\<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub> \<otimes> \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub> = \<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub>" 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 | 134 | moreover have "\<one>\<^bsub>R\<lparr>carrier := H\<rparr>\<^esub> \<otimes> \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub> = \<one>\<^bsub>R\<lparr>carrier := K\<rparr>\<^esub>" 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 | 135 | ultimately show ?case using assms generate_field.one by metis | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | case (incl h) thus ?case using generate_field.incl by force | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | case (a_inv h) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | note hyp = this | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | have "a_inv (R\<lparr>carrier := K\<rparr>) h = a_inv R h" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | using assms group.m_inv_consistent[of "add_monoid R" K] a_comm_group incl_HK[of K] hyp | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | subring.axioms(1)[OF subfieldE(1)[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 | 144 | unfolding comm_group_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 | 145 | moreover have "a_inv (R\<lparr>carrier := H\<rparr>) h = a_inv R h" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | using assms group.m_inv_consistent[of "add_monoid R" H] a_comm_group incl_HK[of H] hyp | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 | subring.axioms(1)[OF subfieldE(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 | 148 | unfolding comm_group_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 | 149 | ultimately show ?case using generate_field.a_inv a_inv.IH 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 | 150 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | case (m_inv h) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | note hyp = this | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 |       have h_K : "h \<in> (K - {\<zero>})" using incl_HK[OF assms(2) assms(4)] hyp by auto
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | hence "m_inv (R\<lparr>carrier := K\<rparr>) h = m_inv R h" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | using field.m_inv_mult_of[OF subfield_iff(2)[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 | 156 |                group.m_inv_consistent[of "mult_of R" "K - {\<zero>}"] field_mult_group units_of_inv
 | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | subgroup_mult_of subfieldE[OF assms(2)] unfolding mult_of_def 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 | 158 | by (metis h_K mult_of_def mult_of_is_Units subgroup.mem_carrier units_of_carrier 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 | 159 |       moreover have h_H : "h \<in> (H - {\<zero>})" using incl_HK[OF assms(1) assms(3)] hyp by auto
 | 
| 
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 "m_inv (R\<lparr>carrier := H\<rparr>) h = m_inv R h" | 
| 
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 field.m_inv_mult_of[OF subfield_iff(2)[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 | 162 |                group.m_inv_consistent[of "mult_of R" "H - {\<zero>}"] 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 | 163 | subgroup_mult_of[OF assms(1)] unfolding mult_of_def 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 | 164 | by (metis h_H field_Units m_inv_mult_of mult_of_is_Units subgroup.mem_carrier units_of_def) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | ultimately show ?case using generate_field.m_inv m_inv.IH h_H 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 | 166 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | case (eng_add h1 h2) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | thus ?case using incl_HK assms generate_field.eng_add by force | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | next | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | case (eng_mult h1 h2) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | thus ?case using generate_field.eng_mult by force | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | qed} | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | thus "\<And>x. x \<in> generate_field (R\<lparr>carrier := K\<rparr>) I \<Longrightarrow> x \<in> generate_field (R\<lparr>carrier := H\<rparr>) I" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | by auto | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | lemma (in field) subfield_gen_equality: | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | assumes "subfield H R" "K \<subseteq> H" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | shows "generate_field R K = generate_field (R \<lparr> carrier := H \<rparr>) K" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | using subfield_gen_incl[OF assms(1) carrier_is_subfield assms(2)] assms subringE(1) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | subfield_gen_incl[OF carrier_is_subfield assms(1) _ assms(2)] subfieldE(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 | 182 | by force | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | end |