src/HOL/Algebra/Ideal.thy
author wenzelm
Sun, 10 Mar 2019 23:23:03 +0100
changeset 69895 6b03a8cf092d
parent 69712 dc85b5b3a532
child 70215 8371a25ca177
permissions -rw-r--r--
more formal contributors (with the help of the history);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 41433
diff changeset
     1
(*  Title:      HOL/Algebra/Ideal.thy
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     2
    Author:     Stephan Hohe, TU Muenchen
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     3
*)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     4
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     5
theory Ideal
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     6
imports Ring AbelCoset
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     7
begin
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     8
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
     9
section \<open>Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    10
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
    11
subsection \<open>Definitions\<close>
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    12
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
    13
subsubsection \<open>General definition\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    14
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
    15
locale ideal = additive_subgroup I R + ring R for I and R (structure) +
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    16
  assumes I_l_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
    17
      and I_r_closed: "\<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    18
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
    19
sublocale ideal \<subseteq> abelian_subgroup I R
68458
023b353911c5 Algebra tidy-up
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
    20
proof (intro abelian_subgroupI3 abelian_group.intro)
023b353911c5 Algebra tidy-up
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
    21
  show "additive_subgroup I R"
023b353911c5 Algebra tidy-up
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
    22
    by (simp add: is_additive_subgroup)
023b353911c5 Algebra tidy-up
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
    23
  show "abelian_monoid R"
023b353911c5 Algebra tidy-up
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
    24
    by (simp add: abelian_monoid_axioms)
023b353911c5 Algebra tidy-up
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
    25
  show "abelian_group_axioms R"
023b353911c5 Algebra tidy-up
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
    26
    using abelian_group_def is_abelian_group by blast
023b353911c5 Algebra tidy-up
paulson <lp15@cam.ac.uk>
parents: 68445
diff changeset
    27
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    28
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    29
lemma (in ideal) is_ideal: "ideal I R"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    30
  by (rule ideal_axioms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    31
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    32
lemma idealI:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
    33
  fixes R (structure)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
    34
  assumes "ring R"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
    35
  assumes a_subgroup: "subgroup I (add_monoid R)"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    36
    and I_l_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    37
    and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    38
  shows "ideal I R"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
    39
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
    40
  interpret ring R by fact
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
    41
  show ?thesis  
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    42
    by (auto simp: ideal.intro ideal_axioms.intro additive_subgroupI a_subgroup ring_axioms I_l_closed I_r_closed)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
    43
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    44
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
    45
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
    46
subsubsection (in ring) \<open>Ideals Generated by a Subset of \<^term>\<open>carrier R\<close>\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    47
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    48
definition genideal :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set"  ("Idl\<index> _" [80] 79)
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61506
diff changeset
    49
  where "genideal R S = \<Inter>{I. ideal I R \<and> S \<subseteq> I}"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    50
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
    51
subsubsection \<open>Principal Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    52
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    53
locale principalideal = ideal +
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    54
  assumes generate: "\<exists>i \<in> carrier R. I = Idl {i}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    55
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    56
lemma (in principalideal) is_principalideal: "principalideal I R"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    57
  by (rule principalideal_axioms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    58
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    59
lemma principalidealI:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
    60
  fixes R (structure)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
    61
  assumes "ideal I R"
47409
c5be1120980d tuned proofs;
wenzelm
parents: 44677
diff changeset
    62
    and generate: "\<exists>i \<in> carrier R. I = Idl {i}"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    63
  shows "principalideal I R"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
    64
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
    65
  interpret ideal I R by fact
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    66
  show ?thesis
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    67
    by (intro principalideal.intro principalideal_axioms.intro)
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
    68
      (rule is_ideal, rule generate)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
    69
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    70
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    71
(* NEW ====== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    72
lemma (in ideal) rcos_const_imp_mem:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    73
  assumes "i \<in> carrier R" and "I +> i = I" shows "i \<in> I"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    74
  using additive_subgroup.zero_closed[OF ideal.axioms(1)[OF ideal_axioms]] assms
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    75
  by (force simp add: a_r_coset_def')
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    76
(* ========== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    77
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    78
(* NEW ====== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    79
lemma (in ring) a_rcos_zero:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    80
  assumes "ideal I R" "i \<in> I" shows "I +> i = I"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    81
  using abelian_subgroupI3[OF ideal.axioms(1) is_abelian_group]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    82
  by (simp add: abelian_subgroup.a_rcos_const assms)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    83
(* ========== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    84
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    85
(* NEW ====== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    86
lemma (in ring) ideal_is_normal:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    87
  assumes "ideal I R" shows "I \<lhd> (add_monoid R)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    88
  using abelian_subgroup.a_normal[OF abelian_subgroupI3[OF ideal.axioms(1)]]
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    89
        abelian_group_axioms assms
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    90
  by auto 
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    91
(* ========== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    92
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    93
(* NEW ====== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    94
lemma (in ideal) a_rcos_sum:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    95
  assumes "a \<in> carrier R" and "b \<in> carrier R" shows "(I +> a) <+> (I +> b) = I +> (a \<oplus> b)"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    96
  using normal.rcos_sum[OF ideal_is_normal[OF ideal_axioms]] assms
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    97
  unfolding set_add_def a_r_coset_def by simp
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    98
(* ========== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
    99
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   100
(* NEW ====== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   101
lemma (in ring) set_add_comm:
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   102
  assumes "I \<subseteq> carrier R" "J \<subseteq> carrier R" shows "I <+> J = J <+> I"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   103
proof -
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   104
  { fix I J assume "I \<subseteq> carrier R" "J \<subseteq> carrier R" hence "I <+> J \<subseteq> J <+> I"
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   105
      using a_comm unfolding set_add_def' by (auto, blast) }
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   106
  thus ?thesis
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   107
    using assms by auto
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   108
qed
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   109
(* ========== *)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   110
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   111
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   112
subsubsection \<open>Maximal Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   113
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   114
locale maximalideal = ideal +
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   115
  assumes I_notcarr: "carrier R \<noteq> I"
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   116
    and I_maximal: "\<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> (J = I) \<or> (J = carrier R)"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   117
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   118
lemma (in maximalideal) is_maximalideal: "maximalideal I R"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   119
  by (rule maximalideal_axioms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   120
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   121
lemma maximalidealI:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   122
  fixes R
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   123
  assumes "ideal I R"
47409
c5be1120980d tuned proofs;
wenzelm
parents: 44677
diff changeset
   124
    and I_notcarr: "carrier R \<noteq> I"
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   125
    and I_maximal: "\<And>J. \<lbrakk>ideal J R; I \<subseteq> J; J \<subseteq> carrier R\<rbrakk> \<Longrightarrow> (J = I) \<or> (J = carrier R)"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   126
  shows "maximalideal I R"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   127
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   128
  interpret ideal I R by fact
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   129
  show ?thesis
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   130
    by (intro maximalideal.intro maximalideal_axioms.intro)
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   131
      (rule is_ideal, rule I_notcarr, rule I_maximal)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   132
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   133
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   134
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   135
subsubsection \<open>Prime Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   136
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   137
locale primeideal = ideal + cring +
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   138
  assumes I_notcarr: "carrier R \<noteq> I"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   139
    and I_prime: "\<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   140
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   141
lemma (in primeideal) primeideal: "primeideal I R"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   142
  by (rule primeideal_axioms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   143
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   144
lemma primeidealI:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   145
  fixes R (structure)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   146
  assumes "ideal I R"
47409
c5be1120980d tuned proofs;
wenzelm
parents: 44677
diff changeset
   147
    and "cring R"
c5be1120980d tuned proofs;
wenzelm
parents: 44677
diff changeset
   148
    and I_notcarr: "carrier R \<noteq> I"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   149
    and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   150
  shows "primeideal I R"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   151
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   152
  interpret ideal I R by fact
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   153
  interpret cring R by fact
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   154
  show ?thesis
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   155
    by (intro primeideal.intro primeideal_axioms.intro)
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   156
      (rule is_ideal, rule is_cring, rule I_notcarr, rule I_prime)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   157
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   158
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   159
lemma primeidealI2:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   160
  fixes R (structure)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   161
  assumes "additive_subgroup I R"
47409
c5be1120980d tuned proofs;
wenzelm
parents: 44677
diff changeset
   162
    and "cring R"
c5be1120980d tuned proofs;
wenzelm
parents: 44677
diff changeset
   163
    and I_l_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> x \<otimes> a \<in> I"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   164
    and I_r_closed: "\<And>a x. \<lbrakk>a \<in> I; x \<in> carrier R\<rbrakk> \<Longrightarrow> a \<otimes> x \<in> I"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   165
    and I_notcarr: "carrier R \<noteq> I"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   166
    and I_prime: "\<And>a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   167
  shows "primeideal I R"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   168
proof -
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   169
  interpret additive_subgroup I R by fact
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   170
  interpret cring R by fact
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   171
  show ?thesis apply intro_locales
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   172
    apply (intro ideal_axioms.intro)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   173
    apply (erule (1) I_l_closed)
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   174
    apply (erule (1) I_r_closed)
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   175
    by (simp add: I_notcarr I_prime primeideal_axioms.intro)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   176
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   177
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   178
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   179
subsection \<open>Special Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   180
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   181
lemma (in ring) zeroideal: "ideal {\<zero>} R"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   182
  by (intro idealI subgroup.intro) (simp_all add: ring_axioms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   183
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   184
lemma (in ring) oneideal: "ideal (carrier R) R"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   185
  by (rule idealI) (auto intro: ring_axioms add.subgroupI)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   186
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   187
lemma (in "domain") zeroprimeideal: "primeideal {\<zero>} R"
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   188
proof -
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   189
  have "carrier R \<noteq> {\<zero>}"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   190
    by (simp add: carrier_one_not_zero)
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   191
  then show ?thesis
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   192
    by (metis (no_types, lifting) domain_axioms domain_def integral primeidealI singleton_iff zeroideal)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   193
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   194
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   195
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   196
subsection \<open>General Ideal Properies\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   197
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   198
lemma (in ideal) one_imp_carrier:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   199
  assumes I_one_closed: "\<one> \<in> I"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   200
  shows "I = carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   201
proof
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   202
  show "carrier R \<subseteq> I"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   203
    using I_r_closed assms by fastforce
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   204
  show "I \<subseteq> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   205
    by (rule a_subset)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   206
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   207
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   208
lemma (in ideal) Icarr:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   209
  assumes iI: "i \<in> I"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   210
  shows "i \<in> carrier R"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   211
  using iI by (rule a_Hcarr)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   212
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   213
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   214
subsection \<open>Intersection of Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   215
61506
wenzelm
parents: 61382
diff changeset
   216
paragraph \<open>Intersection of two ideals\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
   217
text \<open>The intersection of any two ideals is again an ideal in \<^term>\<open>R\<close>\<close>
61506
wenzelm
parents: 61382
diff changeset
   218
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   219
lemma (in ring) i_intersect:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   220
  assumes "ideal I R"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   221
  assumes "ideal J R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   222
  shows "ideal (I \<inter> J) R"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   223
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   224
  interpret ideal I R by fact
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   225
  interpret ideal J R by fact
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   226
  have IJ: "I \<inter> J \<subseteq> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   227
    by (force simp: a_subset)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   228
  show ?thesis
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   229
    apply (intro idealI subgroup.intro)
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   230
    apply (simp_all add: IJ ring_axioms I_l_closed assms ideal.I_l_closed ideal.I_r_closed flip: a_inv_def)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   231
    done
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   232
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   233
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
   234
text \<open>The intersection of any Number of Ideals is again an Ideal in \<^term>\<open>R\<close>\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   235
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   236
lemma (in ring) i_Intersect:
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   237
  assumes Sideals: "\<And>I. I \<in> S \<Longrightarrow> ideal I R" and notempty: "S \<noteq> {}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   238
  shows "ideal (\<Inter>S) R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   239
proof -
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   240
  { fix x y J
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   241
    assume "\<forall>I\<in>S. x \<in> I" "\<forall>I\<in>S. y \<in> I" and JS: "J \<in> S"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   242
    interpret ideal J R by (rule Sideals[OF JS])
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   243
    have "x \<oplus> y \<in> J"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   244
      by (simp add: JS \<open>\<forall>I\<in>S. x \<in> I\<close> \<open>\<forall>I\<in>S. y \<in> I\<close>) }
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   245
  moreover
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   246
    have "\<zero> \<in> J" if "J \<in> S" for J
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   247
      by (simp add: that Sideals additive_subgroup.zero_closed ideal.axioms(1)) 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   248
  moreover
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   249
  { fix x J
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   250
    assume "\<forall>I\<in>S. x \<in> I" and JS: "J \<in> S"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   251
    interpret ideal J R by (rule Sideals[OF JS])
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   252
    have "\<ominus> x \<in> J"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   253
      by (simp add: JS \<open>\<forall>I\<in>S. x \<in> I\<close>) }
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   254
  moreover
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   255
  { fix x y J
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   256
    assume "\<forall>I\<in>S. x \<in> I" and ycarr: "y \<in> carrier R" and JS: "J \<in> S"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   257
    interpret ideal J R by (rule Sideals[OF JS])
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   258
    have "y \<otimes> x \<in> J" "x \<otimes> y \<in> J" 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   259
      using I_l_closed I_r_closed JS \<open>\<forall>I\<in>S. x \<in> I\<close> ycarr by blast+ }
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   260
  moreover
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   261
  { fix x
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   262
    assume "\<forall>I\<in>S. x \<in> I"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   263
    obtain I0 where I0S: "I0 \<in> S"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   264
      using notempty by blast
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   265
    interpret ideal I0 R by (rule Sideals[OF I0S])
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   266
    have "x \<in> I0"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   267
      by (simp add: I0S \<open>\<forall>I\<in>S. x \<in> I\<close>) 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   268
    with a_subset have "x \<in> carrier R" by fast }
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   269
  ultimately show ?thesis
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   270
    by unfold_locales (auto simp: Inter_eq simp flip: a_inv_def)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   271
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   272
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   273
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   274
subsection \<open>Addition of Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   275
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   276
lemma (in ring) add_ideals:
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   277
  assumes idealI: "ideal I R" and idealJ: "ideal J R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   278
  shows "ideal (I <+> J) R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   279
proof (rule ideal.intro)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   280
  show "additive_subgroup (I <+> J) R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   281
    by (intro ideal.axioms[OF idealI] ideal.axioms[OF idealJ] add_additive_subgroups)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   282
  show "ring R"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   283
    by (rule ring_axioms)
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   284
  show "ideal_axioms (I <+> J) R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   285
  proof -
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   286
    { fix x i j
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   287
      assume xcarr: "x \<in> carrier R" and iI: "i \<in> I" and jJ: "j \<in> J"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   288
      from xcarr ideal.Icarr[OF idealI iI] ideal.Icarr[OF idealJ jJ]
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   289
      have "\<exists>h\<in>I. \<exists>k\<in>J. (i \<oplus> j) \<otimes> x = h \<oplus> k"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   290
        by (meson iI ideal.I_r_closed idealJ jJ l_distr local.idealI) }
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   291
    moreover
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   292
    { fix x i j
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   293
      assume xcarr: "x \<in> carrier R" and iI: "i \<in> I" and jJ: "j \<in> J"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   294
      from xcarr ideal.Icarr[OF idealI iI] ideal.Icarr[OF idealJ jJ]
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   295
      have "\<exists>h\<in>I. \<exists>k\<in>J. x \<otimes> (i \<oplus> j) = h \<oplus> k"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   296
        by (meson iI ideal.I_l_closed idealJ jJ local.idealI r_distr) }
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   297
    ultimately show "ideal_axioms (I <+> J) R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   298
      by (intro ideal_axioms.intro) (auto simp: set_add_defs)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   299
  qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   300
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   301
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
   302
subsection (in ring) \<open>Ideals generated by a subset of \<^term>\<open>carrier R\<close>\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   303
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
   304
text \<open>\<^term>\<open>genideal\<close> generates an ideal\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   305
lemma (in ring) genideal_ideal:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   306
  assumes Scarr: "S \<subseteq> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   307
  shows "ideal (Idl S) R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   308
unfolding genideal_def
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   309
proof (rule i_Intersect, fast, simp)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   310
  from oneideal and Scarr
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   311
  show "\<exists>I. ideal I R \<and> S \<le> I" by fast
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   312
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   313
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   314
lemma (in ring) genideal_self:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   315
  assumes "S \<subseteq> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   316
  shows "S \<subseteq> Idl S"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   317
  unfolding genideal_def by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   319
lemma (in ring) genideal_self':
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   320
  assumes carr: "i \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   321
  shows "i \<in> Idl {i}"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   322
  by (simp add: genideal_def)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   323
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
   324
text \<open>\<^term>\<open>genideal\<close> generates the minimal ideal\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   325
lemma (in ring) genideal_minimal:
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   326
  assumes "ideal I R" "S \<subseteq> I"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   327
  shows "Idl S \<subseteq> I"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   328
  unfolding genideal_def by rule (elim InterD, simp add: assms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   329
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   330
text \<open>Generated ideals and subsets\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   331
lemma (in ring) Idl_subset_ideal:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   332
  assumes Iideal: "ideal I R"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   333
    and Hcarr: "H \<subseteq> carrier R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   334
  shows "(Idl H \<subseteq> I) = (H \<subseteq> I)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   335
proof
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   336
  assume a: "Idl H \<subseteq> I"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 20318
diff changeset
   337
  from Hcarr have "H \<subseteq> Idl H" by (rule genideal_self)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   338
  with a show "H \<subseteq> I" by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   339
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   340
  fix x
47409
c5be1120980d tuned proofs;
wenzelm
parents: 44677
diff changeset
   341
  assume "H \<subseteq> I"
c5be1120980d tuned proofs;
wenzelm
parents: 44677
diff changeset
   342
  with Iideal have "I \<in> {I. ideal I R \<and> H \<subseteq> I}" by fast
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   343
  then show "Idl H \<subseteq> I" unfolding genideal_def by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   344
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   345
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   346
lemma (in ring) subset_Idl_subset:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   347
  assumes Icarr: "I \<subseteq> carrier R"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   348
    and HI: "H \<subseteq> I"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   349
  shows "Idl H \<subseteq> Idl I"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   350
proof -
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   351
  from Icarr have Iideal: "ideal (Idl I) R"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   352
    by (rule genideal_ideal)
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   353
  from HI and Icarr have "H \<subseteq> carrier R"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   354
    by fast
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   355
  with Iideal have "(H \<subseteq> Idl I) = (Idl H \<subseteq> Idl I)"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   356
    by (rule Idl_subset_ideal[symmetric])
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   357
  then show "Idl H \<subseteq> Idl I"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   358
    by (meson HI Icarr genideal_self order_trans)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   359
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   360
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   361
lemma (in ring) Idl_subset_ideal':
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   362
  assumes acarr: "a \<in> carrier R" and bcarr: "b \<in> carrier R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   363
  shows "Idl {a} \<subseteq> Idl {b} \<longleftrightarrow> a \<in> Idl {b}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   364
proof -
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   365
  have "Idl {a} \<subseteq> Idl {b} \<longleftrightarrow> {a} \<subseteq> Idl {b}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   366
    by (simp add: Idl_subset_ideal acarr bcarr genideal_ideal)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   367
  also have "\<dots> \<longleftrightarrow> a \<in> Idl {b}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   368
    by blast
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   369
  finally show ?thesis .
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   370
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   371
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   372
lemma (in ring) genideal_zero: "Idl {\<zero>} = {\<zero>}"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   373
proof
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   374
  show "Idl {\<zero>} \<subseteq> {\<zero>}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   375
    by (simp add: genideal_minimal zeroideal)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   376
  show "{\<zero>} \<subseteq> Idl {\<zero>}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   377
    by (simp add: genideal_self')
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   378
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   379
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   380
lemma (in ring) genideal_one: "Idl {\<one>} = carrier R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   381
proof -
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   382
  interpret ideal "Idl {\<one>}" "R" by (rule genideal_ideal) fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   383
  show "Idl {\<one>} = carrier R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   384
    using genideal_self' one_imp_carrier by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   385
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   386
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   387
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   388
text \<open>Generation of Principal Ideals in Commutative Rings\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   389
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   390
definition cgenideal :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set"  ("PIdl\<index> _" [80] 79)
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   391
  where "cgenideal R a = {x \<otimes>\<^bsub>R\<^esub> a | x. x \<in> carrier R}"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   392
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   393
text \<open>genhideal (?) really generates an ideal\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   394
lemma (in cring) cgenideal_ideal:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   395
  assumes acarr: "a \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   396
  shows "ideal (PIdl a) R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   397
  unfolding cgenideal_def
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   398
proof (intro subgroup.intro idealI[OF ring_axioms], simp_all)
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   399
  show "{x \<otimes> a |x. x \<in> carrier R} \<subseteq> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   400
    by (blast intro: acarr)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   401
  show "\<And>x y. \<lbrakk>\<exists>u. x = u \<otimes> a \<and> u \<in> carrier R; \<exists>x. y = x \<otimes> a \<and> x \<in> carrier R\<rbrakk>
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   402
              \<Longrightarrow> \<exists>v. x \<oplus> y = v \<otimes> a \<and> v \<in> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   403
    by (metis assms cring.cring_simprules(1) is_cring l_distr)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   404
  show "\<exists>x. \<zero> = x \<otimes> a \<and> x \<in> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   405
    by (metis assms l_null zero_closed)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   406
  show "\<And>x. \<exists>u. x = u \<otimes> a \<and> u \<in> carrier R 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   407
            \<Longrightarrow> \<exists>v. inv\<^bsub>add_monoid R\<^esub> x = v \<otimes> a \<and> v \<in> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   408
    by (metis a_inv_def add.inv_closed assms l_minus)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   409
  show "\<And>b x. \<lbrakk>\<exists>x. b = x \<otimes> a \<and> x \<in> carrier R; x \<in> carrier R\<rbrakk>
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   410
       \<Longrightarrow> \<exists>z. x \<otimes> b = z \<otimes> a \<and> z \<in> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   411
    by (metis assms m_assoc m_closed)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   412
  show "\<And>b x. \<lbrakk>\<exists>x. b = x \<otimes> a \<and> x \<in> carrier R; x \<in> carrier R\<rbrakk>
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   413
       \<Longrightarrow> \<exists>z. b \<otimes> x = z \<otimes> a \<and> z \<in> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   414
    by (metis assms m_assoc m_comm m_closed)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   415
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   416
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   417
lemma (in ring) cgenideal_self:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   418
  assumes icarr: "i \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   419
  shows "i \<in> PIdl i"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   420
  unfolding cgenideal_def
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   421
proof simp
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   422
  from icarr have "i = \<one> \<otimes> i"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   423
    by simp
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   424
  with icarr show "\<exists>x. i = x \<otimes> i \<and> x \<in> carrier R"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   425
    by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   426
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   427
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69122
diff changeset
   428
text \<open>\<^const>\<open>cgenideal\<close> is minimal\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   429
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   430
lemma (in ring) cgenideal_minimal:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   431
  assumes "ideal J R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   432
  assumes aJ: "a \<in> J"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   433
  shows "PIdl a \<subseteq> J"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   434
proof -
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   435
  interpret ideal J R by fact
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   436
  show ?thesis
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   437
    unfolding cgenideal_def
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   438
    using I_l_closed aJ by blast
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   439
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   440
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   441
lemma (in cring) cgenideal_eq_genideal:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   442
  assumes icarr: "i \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   443
  shows "PIdl i = Idl {i}"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   444
proof
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   445
  show "PIdl i \<subseteq> Idl {i}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   446
    by (simp add: cgenideal_minimal genideal_ideal genideal_self' icarr)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   447
  show "Idl {i} \<subseteq> PIdl i"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   448
    by (simp add: cgenideal_ideal cgenideal_self genideal_minimal icarr)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   449
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   450
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   451
lemma (in cring) cgenideal_eq_rcos: "PIdl i = carrier R #> i"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   452
  unfolding cgenideal_def r_coset_def by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   453
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   454
lemma (in cring) cgenideal_is_principalideal:
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   455
  assumes "i \<in> carrier R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   456
  shows "principalideal (PIdl i) R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   457
proof -
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   458
  have "\<exists>i'\<in>carrier R. PIdl i = Idl {i'}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   459
    using cgenideal_eq_genideal assms by auto
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   460
  then show ?thesis
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   461
    by (simp add: cgenideal_ideal assms principalidealI)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   462
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   463
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   464
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   465
subsection \<open>Union of Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   466
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   467
lemma (in ring) union_genideal:
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   468
  assumes idealI: "ideal I R" and idealJ: "ideal J R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   469
  shows "Idl (I \<union> J) = I <+> J"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   470
proof
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   471
  show "Idl (I \<union> J) \<subseteq> I <+> J"
69122
1b5178abaf97 updates to Algebra from Baillon and de Vilhena
paulson <lp15@cam.ac.uk>
parents: 68687
diff changeset
   472
  proof (rule ring.genideal_minimal [OF ring_axioms])
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   473
    show "ideal (I <+> J) R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   474
      by (rule add_ideals[OF idealI idealJ])
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   475
    have "\<And>x. x \<in> I \<Longrightarrow> \<exists>xa\<in>I. \<exists>xb\<in>J. x = xa \<oplus> xb"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   476
      by (metis additive_subgroup.zero_closed ideal.Icarr idealJ ideal_def local.idealI r_zero)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   477
    moreover have "\<And>x. x \<in> J \<Longrightarrow> \<exists>xa\<in>I. \<exists>xb\<in>J. x = xa \<oplus> xb"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   478
      by (metis additive_subgroup.zero_closed ideal.Icarr idealJ ideal_def l_zero local.idealI)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   479
    ultimately show "I \<union> J \<subseteq> I <+> J"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   480
      by (auto simp: set_add_defs) 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   481
  qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   482
next
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   483
  show "I <+> J \<subseteq> Idl (I \<union> J)"
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   484
    by (auto simp: set_add_defs genideal_def additive_subgroup.a_closed ideal_def subsetD)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   485
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   486
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   487
subsection \<open>Properties of Principal Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   488
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   489
text \<open>The zero ideal is a principal ideal\<close>
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   490
corollary (in ring) zeropideal: "principalideal {\<zero>} R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   491
  using genideal_zero principalidealI zeroideal by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   492
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   493
text \<open>The unit ideal is a principal ideal\<close>
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   494
corollary (in ring) onepideal: "principalideal (carrier R) R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   495
  using genideal_one oneideal principalidealI by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   496
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   497
text \<open>Every principal ideal is a right coset of the carrier\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   498
lemma (in principalideal) rcos_generate:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   499
  assumes "cring R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   500
  shows "\<exists>x\<in>I. I = carrier R #> x"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   501
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   502
  interpret cring R by fact
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   503
  from generate obtain i where icarr: "i \<in> carrier R" and I1: "I = Idl {i}"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   504
    by fast+
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   505
  then have "I = PIdl i"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   506
    by (simp add: cgenideal_eq_genideal)
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   507
  moreover have "i \<in> I"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   508
    by (simp add: I1 genideal_self' icarr)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   509
  moreover have "PIdl i = carrier R #> i"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   510
    unfolding cgenideal_def r_coset_def by fast
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   511
  ultimately show "\<exists>x\<in>I. I = carrier R #> x"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   512
    by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   513
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   514
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   515
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   516
text \<open>This next lemma would be trivial if placed in a theory that imports QuotRing,
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   517
      but it makes more sense to have it here (easier to find and coherent with the
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   518
      previous developments).\<close>
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   519
69895
6b03a8cf092d more formal contributors (with the help of the history);
wenzelm
parents: 69712
diff changeset
   520
lemma (in cring) cgenideal_prod: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close>
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   521
  assumes "a \<in> carrier R" "b \<in> carrier R"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   522
  shows "(PIdl a) <#> (PIdl b) = PIdl (a \<otimes> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   523
proof -
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   524
  have "(carrier R #> a) <#> (carrier R #> b) = carrier R #> (a \<otimes> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   525
  proof
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   526
    show "(carrier R #> a) <#> (carrier R #> b) \<subseteq> carrier R #> a \<otimes> b"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   527
    proof
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   528
      fix x assume "x \<in> (carrier R #> a) <#> (carrier R #> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   529
      then obtain r1 r2 where r1: "r1 \<in> carrier R" and r2: "r2 \<in> carrier R"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   530
                          and "x = (r1 \<otimes> a) \<otimes> (r2 \<otimes> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   531
        unfolding set_mult_def r_coset_def by blast
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   532
      hence "x = (r1 \<otimes> r2) \<otimes> (a \<otimes> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   533
        by (simp add: assms local.ring_axioms m_lcomm ring.ring_simprules(11))
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   534
      thus "x \<in> carrier R #> a \<otimes> b"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   535
        unfolding r_coset_def using r1 r2 assms by blast 
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   536
    qed
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   537
  next
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   538
    show "carrier R #> a \<otimes> b \<subseteq> (carrier R #> a) <#> (carrier R #> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   539
    proof
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   540
      fix x assume "x \<in> carrier R #> a \<otimes> b"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   541
      then obtain r where r: "r \<in> carrier R" "x = r \<otimes> (a \<otimes> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   542
        unfolding r_coset_def by blast
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   543
      hence "x = (r \<otimes> a) \<otimes> (\<one> \<otimes> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   544
        using assms by (simp add: m_assoc)
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   545
      thus "x \<in> (carrier R #> a) <#> (carrier R #> b)"
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   546
        unfolding set_mult_def r_coset_def using assms r by blast
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   547
    qed
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   548
  qed
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   549
  thus ?thesis
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   550
    using cgenideal_eq_rcos[of a] cgenideal_eq_rcos[of b] cgenideal_eq_rcos[of "a \<otimes> b"] by simp
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   551
qed
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   552
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   553
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   554
subsection \<open>Prime Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   555
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   556
lemma (in ideal) primeidealCD:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   557
  assumes "cring R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   558
  assumes notprime: "\<not> primeideal I R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   559
  shows "carrier R = I \<or> (\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   560
proof (rule ccontr, clarsimp)
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   561
  interpret cring R by fact
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   562
  assume InR: "carrier R \<noteq> I"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   563
    and "\<forall>a. a \<in> carrier R \<longrightarrow> (\<forall>b. a \<otimes> b \<in> I \<longrightarrow> b \<in> carrier R \<longrightarrow> a \<in> I \<or> b \<in> I)"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   564
  then have I_prime: "\<And> a b. \<lbrakk>a \<in> carrier R; b \<in> carrier R; a \<otimes> b \<in> I\<rbrakk> \<Longrightarrow> a \<in> I \<or> b \<in> I"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   565
    by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   566
  have "primeideal I R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   567
    by (simp add: I_prime InR is_cring is_ideal primeidealI)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   568
  with notprime show False by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   569
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   570
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   571
lemma (in ideal) primeidealCE:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   572
  assumes "cring R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   573
  assumes notprime: "\<not> primeideal I R"
23383
5460951833fa tuned proofs: avoid implicit prems;
wenzelm
parents: 23350
diff changeset
   574
  obtains "carrier R = I"
5460951833fa tuned proofs: avoid implicit prems;
wenzelm
parents: 23350
diff changeset
   575
    | "\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   576
proof -
30729
461ee3e49ad3 interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents: 30363
diff changeset
   577
  interpret R: cring R by fact
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   578
  assume "carrier R = I ==> thesis"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   579
    and "\<exists>a b. a \<in> carrier R \<and> b \<in> carrier R \<and> a \<otimes> b \<in> I \<and> a \<notin> I \<and> b \<notin> I \<Longrightarrow> thesis"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   580
  then show thesis using primeidealCD [OF R.is_cring notprime] by blast
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   581
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   582
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 63040
diff changeset
   583
text \<open>If \<open>{\<zero>}\<close> is a prime ideal of a commutative ring, the ring is a domain\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   584
lemma (in cring) zeroprimeideal_domainI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   585
  assumes pi: "primeideal {\<zero>} R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   586
  shows "domain R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   587
proof (intro domain.intro is_cring domain_axioms.intro)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   588
  show "\<one> \<noteq> \<zero>"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   589
    using genideal_one genideal_zero pi primeideal.I_notcarr by force
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   590
  show "a = \<zero> \<or> b = \<zero>" if ab: "a \<otimes> b = \<zero>" and carr: "a \<in> carrier R" "b \<in> carrier R" for a b
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   591
  proof -
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   592
    interpret primeideal "{\<zero>}" "R" by (rule pi)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   593
    show "a = \<zero> \<or> b = \<zero>"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   594
      using I_prime ab carr by blast
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   595
  qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   596
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   597
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   598
corollary (in cring) domain_eq_zeroprimeideal: "domain R = primeideal {\<zero>} R"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   599
  using domain.zeroprimeideal zeroprimeideal_domainI by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   600
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   601
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   602
subsection \<open>Maximal Ideals\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   603
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   604
lemma (in ideal) helper_I_closed:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   605
  assumes carr: "a \<in> carrier R" "x \<in> carrier R" "y \<in> carrier R"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   606
    and axI: "a \<otimes> x \<in> I"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   607
  shows "a \<otimes> (x \<otimes> y) \<in> I"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   608
proof -
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   609
  from axI and carr have "(a \<otimes> x) \<otimes> y \<in> I"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   610
    by (simp add: I_r_closed)
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   611
  also from carr have "(a \<otimes> x) \<otimes> y = a \<otimes> (x \<otimes> y)"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   612
    by (simp add: m_assoc)
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   613
  finally show "a \<otimes> (x \<otimes> y) \<in> I" .
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   614
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   615
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   616
lemma (in ideal) helper_max_prime:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   617
  assumes "cring R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   618
  assumes acarr: "a \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   619
  shows "ideal {x\<in>carrier R. a \<otimes> x \<in> I} R"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   620
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   621
  interpret cring R by fact
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   622
  show ?thesis 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   623
  proof (rule idealI, simp_all)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   624
    show "ring R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   625
      by (simp add: local.ring_axioms)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   626
    show "subgroup {x \<in> carrier R. a \<otimes> x \<in> I} (add_monoid R)"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   627
      by (rule subgroup.intro) (auto simp: r_distr acarr r_minus simp flip: a_inv_def)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   628
    show "\<And>b x. \<lbrakk>b \<in> carrier R \<and> a \<otimes> b \<in> I; x \<in> carrier R\<rbrakk>
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   629
                 \<Longrightarrow> a \<otimes> (x \<otimes> b) \<in> I"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   630
      using acarr helper_I_closed m_comm by auto
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   631
    show "\<And>b x. \<lbrakk>b \<in> carrier R \<and> a \<otimes> b \<in> I; x \<in> carrier R\<rbrakk>
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   632
                \<Longrightarrow> a \<otimes> (b \<otimes> x) \<in> I"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   633
      by (simp add: acarr helper_I_closed)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   634
  qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   635
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   636
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   637
text \<open>In a cring every maximal ideal is prime\<close>
63633
2accfb71e33b is_prime -> prime
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   638
lemma (in cring) maximalideal_prime:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   639
  assumes "maximalideal I R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   640
  shows "primeideal I R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   641
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 27717
diff changeset
   642
  interpret maximalideal I R by fact
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   643
  show ?thesis 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   644
  proof (rule ccontr)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   645
    assume neg: "\<not> primeideal I R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   646
    then obtain a b where acarr: "a \<in> carrier R" and bcarr: "b \<in> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   647
      and abI: "a \<otimes> b \<in> I" and anI: "a \<notin> I" and bnI: "b \<notin> I" 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   648
      using primeidealCE [OF is_cring]
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   649
      by (metis I_notcarr)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61952
diff changeset
   650
    define J where "J = {x\<in>carrier R. a \<otimes> x \<in> I}"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   651
    from is_cring and acarr have idealJ: "ideal J R"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   652
      unfolding J_def by (rule helper_max_prime)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   653
    have IsubJ: "I \<subseteq> J"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   654
      using I_l_closed J_def a_Hcarr acarr by blast
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   655
    from abI and acarr bcarr have "b \<in> J"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   656
      unfolding J_def by fast
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   657
    with bnI have JnI: "J \<noteq> I" by fast
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   658
    have "\<one> \<notin> J"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   659
      unfolding J_def by (simp add: acarr anI)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   660
    then have Jncarr: "J \<noteq> carrier R" by fast
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   661
    interpret ideal J R by (rule idealJ)    
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26203
diff changeset
   662
    have "J = I \<or> J = carrier R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   663
      by (simp add: I_maximal IsubJ a_subset is_ideal)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   664
    with JnI and Jncarr show False by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   665
  qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   666
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   667
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   668
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   669
subsection \<open>Derived Theorems\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   670
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   671
text \<open>A non-zero cring that has only the two trivial ideals is a field\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   672
lemma (in cring) trivialideals_fieldI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   673
  assumes carrnzero: "carrier R \<noteq> {\<zero>}"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   674
    and haveideals: "{I. ideal I R} = {{\<zero>}, carrier R}"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   675
  shows "field R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   676
proof (intro cring_fieldI equalityI)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   677
  show "Units R \<subseteq> carrier R - {\<zero>}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   678
    by (metis Diff_empty Units_closed Units_r_inv_ex carrnzero l_null one_zeroD subsetI subset_Diff_insert)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   679
  show "carrier R - {\<zero>} \<subseteq> Units R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   680
  proof
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   681
    fix x
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   682
    assume xcarr': "x \<in> carrier R - {\<zero>}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   683
    then have xcarr: "x \<in> carrier R" and xnZ: "x \<noteq> \<zero>" by auto
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   684
    from xcarr have xIdl: "ideal (PIdl x) R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   685
      by (intro cgenideal_ideal) fast
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   686
    have "PIdl x \<noteq> {\<zero>}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   687
      using xcarr xnZ cgenideal_self by blast 
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   688
    with haveideals have "PIdl x = carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   689
      by (blast intro!: xIdl)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   690
    then have "\<one> \<in> PIdl x" by simp
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   691
    then have "\<exists>y. \<one> = y \<otimes> x \<and> y \<in> carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   692
      unfolding cgenideal_def by blast
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   693
    then obtain y where ycarr: " y \<in> carrier R" and ylinv: "\<one> = y \<otimes> x"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   694
      by fast    
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   695
    have "\<exists>y \<in> carrier R. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   696
      using m_comm xcarr ycarr ylinv by auto
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   697
    with xcarr show "x \<in> Units R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   698
      unfolding Units_def by fast
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   699
  qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   700
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   701
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   702
lemma (in field) all_ideals: "{I. ideal I R} = {{\<zero>}, carrier R}"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   703
proof (intro equalityI subsetI)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   704
  fix I
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   705
  assume a: "I \<in> {I. ideal I R}"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   706
  then interpret ideal I R by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   707
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   708
  show "I \<in> {{\<zero>}, carrier R}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   709
  proof (cases "\<exists>a. a \<in> I - {\<zero>}")
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   710
    case True
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   711
    then obtain a where aI: "a \<in> I" and anZ: "a \<noteq> \<zero>"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   712
      by fast+
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   713
    have aUnit: "a \<in> Units R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   714
      by (simp add: aI anZ field_Units)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   715
    then have a: "a \<otimes> inv a = \<one>" by (rule Units_r_inv)
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   716
    from aI and aUnit have "a \<otimes> inv a \<in> I"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   717
      by (simp add: I_r_closed del: Units_r_inv)
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   718
    then have oneI: "\<one> \<in> I" by (simp add: a[symmetric])
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   719
    have "carrier R \<subseteq> I"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   720
      using oneI one_imp_carrier by auto
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   721
    with a_subset have "I = carrier R" by fast
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   722
    then show "I \<in> {{\<zero>}, carrier R}" by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   723
  next
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   724
    case False
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   725
    then have IZ: "\<And>a. a \<in> I \<Longrightarrow> a = \<zero>" by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   726
    have a: "I \<subseteq> {\<zero>}"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   727
      using False by auto
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   728
    have "\<zero> \<in> I" by simp
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   729
    with a have "I = {\<zero>}" by fast
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   730
    then show "I \<in> {{\<zero>}, carrier R}" by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   731
  qed
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   732
qed (auto simp: zeroideal oneideal)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   733
68445
c183a6a69f2d reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents: 68443
diff changeset
   734
\<comment>\<open>"Jacobson Theorem 2.2"\<close>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   735
lemma (in cring) trivialideals_eq_field:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   736
  assumes carrnzero: "carrier R \<noteq> {\<zero>}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   737
  shows "({I. ideal I R} = {{\<zero>}, carrier R}) = field R"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   738
  by (fast intro!: trivialideals_fieldI[OF carrnzero] field.all_ideals)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   739
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   740
61382
efac889fccbc isabelle update_cartouches;
wenzelm
parents: 47409
diff changeset
   741
text \<open>Like zeroprimeideal for domains\<close>
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   742
lemma (in field) zeromaximalideal: "maximalideal {\<zero>} R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   743
proof (intro maximalidealI zeroideal)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   744
  from one_not_zero have "\<one> \<notin> {\<zero>}" by simp
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   745
  with one_closed show "carrier R \<noteq> {\<zero>}" by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   746
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   747
  fix J
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   748
  assume Jideal: "ideal J R"
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   749
  then have "J \<in> {I. ideal I R}" by fast
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   750
  with all_ideals show "J = {\<zero>} \<or> J = carrier R"
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   751
    by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   752
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   753
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   754
lemma (in cring) zeromaximalideal_fieldI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   755
  assumes zeromax: "maximalideal {\<zero>} R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   756
  shows "field R"
68464
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   757
proof (intro trivialideals_fieldI maximalideal.I_notcarr[OF zeromax])
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   758
  have "J = carrier R" if Jn0: "J \<noteq> {\<zero>}" and idealJ: "ideal J R" for J
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   759
  proof -
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   760
    interpret ideal J R by (rule idealJ)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   761
    have "{\<zero>} \<subseteq> J"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   762
      by force
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   763
    from zeromax idealJ this a_subset
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   764
    have "J = {\<zero>} \<or> J = carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   765
      by (rule maximalideal.I_maximal)
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   766
    with Jn0 show "J = carrier R"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   767
      by simp
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   768
  qed
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   769
  then show "{I. ideal I R} = {{\<zero>}, carrier R}"
3ead36cbe6b7 De-applied Ideal.thy
paulson <lp15@cam.ac.uk>
parents: 68458
diff changeset
   770
    by (auto simp: zeroideal oneideal)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   771
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   772
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 44106
diff changeset
   773
lemma (in cring) zeromaximalideal_eq_field: "maximalideal {\<zero>} R = field R"
68443
43055b016688 New material from Martin Baillon and Paulo Emílio de Vilhena
paulson <lp15@cam.ac.uk>
parents: 67443
diff changeset
   774
  using field.zeromaximalideal zeromaximalideal_fieldI by blast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   775
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   776
end