src/HOL/Algebra/Ring.thy
author wenzelm
Sat, 23 May 2015 17:19:37 +0200
changeset 60299 5ae2a2e74c93
parent 60112 3eab4acaa035
child 60773 d09c66a0ea10
permissions -rw-r--r--
clarified NEWS: document_files are officially required since Isabelle2014, but the absence was tolerated as legacy feature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 41433
diff changeset
     1
(*  Title:      HOL/Algebra/Ring.thy
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     2
    Author:     Clemens Ballarin, started 9 December 1996
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
     3
    Copyright:  Clemens Ballarin
20318
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
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27933
diff changeset
     6
theory Ring
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27933
diff changeset
     7
imports FiniteProduct
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
     8
begin
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
     9
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    10
section {* The Algebraic Hierarchy of Rings *}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    11
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    12
subsection {* Abelian Groups *}
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    13
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    14
record 'a ring = "'a monoid" +
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    15
  zero :: 'a ("\<zero>\<index>")
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    16
  add :: "['a, 'a] => 'a" (infixl "\<oplus>\<index>" 65)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    17
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    18
text {* Derived operations. *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    19
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
    20
definition
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    21
  a_inv :: "[('a, 'm) ring_scheme, 'a ] => 'a" ("\<ominus>\<index> _" [81] 80)
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
    22
  where "a_inv R = m_inv \<lparr>carrier = carrier R, mult = add R, one = zero R\<rparr>"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    23
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
    24
definition
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    25
  a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" (infixl "\<ominus>\<index>" 65)
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
    26
  where "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    27
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    28
locale abelian_monoid =
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    29
  fixes G (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    30
  assumes a_comm_monoid:
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
    31
     "comm_monoid \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    32
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    33
definition
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    34
  finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b" where
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
    35
  "finsum G = finprod \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    36
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    37
syntax
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    38
  "_finsum" :: "index => idt => 'a set => 'b => 'b"
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    39
      ("(3\<Oplus>__:_. _)" [1000, 0, 51, 10] 10)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    40
syntax (xsymbols)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    41
  "_finsum" :: "index => idt => 'a set => 'b => 'b"
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    42
      ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    43
syntax (HTML output)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    44
  "_finsum" :: "index => idt => 'a set => 'b => 'b"
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    45
      ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    46
translations
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    47
  "\<Oplus>\<index>i:A. b" == "CONST finsum \<struct>\<index> (%i. b) A"
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    48
  -- {* Beware of argument permutation! *}
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
    49
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    50
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    51
locale abelian_group = abelian_monoid +
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    52
  assumes a_comm_group:
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
    53
     "comm_group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    54
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    55
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    56
subsection {* Basic Properties *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    57
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    58
lemma abelian_monoidI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    59
  fixes R (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    60
  assumes a_closed:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    61
      "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    62
    and zero_closed: "\<zero> \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    63
    and a_assoc:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    64
      "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    65
      (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    66
    and l_zero: "!!x. x \<in> carrier R ==> \<zero> \<oplus> x = x"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    67
    and a_comm:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    68
      "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y = y \<oplus> x"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    69
  shows "abelian_monoid R"
27714
27b4d7c01f8b Tuned (for the sake of a meaningless log entry).
ballarin
parents: 27699
diff changeset
    70
  by (auto intro!: abelian_monoid.intro comm_monoidI intro: assms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    71
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    72
lemma abelian_groupI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    73
  fixes R (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    74
  assumes a_closed:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    75
      "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    76
    and zero_closed: "zero R \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    77
    and a_assoc:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    78
      "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    79
      (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    80
    and a_comm:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    81
      "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==> x \<oplus> y = y \<oplus> x"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    82
    and l_zero: "!!x. x \<in> carrier R ==> \<zero> \<oplus> x = x"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    83
    and l_inv_ex: "!!x. x \<in> carrier R ==> EX y : carrier R. y \<oplus> x = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    84
  shows "abelian_group R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    85
  by (auto intro!: abelian_group.intro abelian_monoidI
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    86
      abelian_group_axioms.intro comm_monoidI comm_groupI
27714
27b4d7c01f8b Tuned (for the sake of a meaningless log entry).
ballarin
parents: 27699
diff changeset
    87
    intro: assms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    88
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    89
lemma (in abelian_monoid) a_monoid:
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
    90
  "monoid \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    91
by (rule comm_monoid.axioms, rule a_comm_monoid) 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    92
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    93
lemma (in abelian_group) a_group:
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
    94
  "group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    95
  by (simp add: group_def a_monoid)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    96
    (simp add: comm_group.axioms group.axioms a_comm_group)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    97
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    98
lemmas monoid_record_simps = partial_object.simps monoid.simps
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
    99
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   100
text {* Transfer facts from multiplicative structures via interpretation. *}
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   101
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   102
sublocale abelian_monoid <
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   103
  add!: monoid "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   104
  where "carrier \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = carrier G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   105
    and "mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = add G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   106
    and "one \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = zero G"
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   107
  by (rule a_monoid) auto
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   108
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   109
context abelian_monoid begin
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   110
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   111
lemmas a_closed = add.m_closed 
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   112
lemmas zero_closed = add.one_closed
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   113
lemmas a_assoc = add.m_assoc
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   114
lemmas l_zero = add.l_one
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   115
lemmas r_zero = add.r_one
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   116
lemmas minus_unique = add.inv_unique
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   117
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   118
end
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   119
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   120
sublocale abelian_monoid <
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   121
  add!: comm_monoid "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   122
  where "carrier \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = carrier G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   123
    and "mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = add G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   124
    and "one \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = zero G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   125
    and "finprod \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = finsum G"
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   126
  by (rule a_comm_monoid) (auto simp: finsum_def)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   127
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   128
context abelian_monoid begin
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   129
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   130
lemmas a_comm = add.m_comm
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   131
lemmas a_lcomm = add.m_lcomm
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   132
lemmas a_ac = a_assoc a_comm a_lcomm
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   133
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   134
lemmas finsum_empty = add.finprod_empty
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   135
lemmas finsum_insert = add.finprod_insert
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   136
lemmas finsum_zero = add.finprod_one
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   137
lemmas finsum_closed = add.finprod_closed
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   138
lemmas finsum_Un_Int = add.finprod_Un_Int
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   139
lemmas finsum_Un_disjoint = add.finprod_Un_disjoint
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   140
lemmas finsum_addf = add.finprod_multf
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   141
lemmas finsum_cong' = add.finprod_cong'
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   142
lemmas finsum_0 = add.finprod_0
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   143
lemmas finsum_Suc = add.finprod_Suc
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   144
lemmas finsum_Suc2 = add.finprod_Suc2
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   145
lemmas finsum_add = add.finprod_mult
60112
3eab4acaa035 finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 59851
diff changeset
   146
lemmas finsum_infinite = add.finprod_infinite
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   147
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   148
lemmas finsum_cong = add.finprod_cong
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   149
text {*Usually, if this rule causes a failed congruence proof error,
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   150
   the reason is that the premise @{text "g \<in> B -> carrier G"} cannot be shown.
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   151
   Adding @{thm [source] Pi_def} to the simpset is often useful. *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   152
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   153
lemmas finsum_reindex = add.finprod_reindex
27699
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   154
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   155
(* The following would be wrong.  Needed is the equivalent of (^) for addition,
27699
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   156
  or indeed the canonical embedding from Nat into the monoid.
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   157
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   158
lemma finsum_const:
27699
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   159
  assumes fin [simp]: "finite A"
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   160
      and a [simp]: "a : carrier G"
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   161
    shows "finsum G (%x. a) A = a (^) card A"
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   162
  using fin apply induct
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   163
  apply force
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   164
  apply (subst finsum_insert)
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   165
  apply auto
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   166
  apply (force simp add: Pi_def)
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   167
  apply (subst m_comm)
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   168
  apply auto
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   169
done
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   170
*)
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   171
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   172
lemmas finsum_singleton = add.finprod_singleton
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   173
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   174
end
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   175
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   176
sublocale abelian_group <
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   177
  add!: group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   178
  where "carrier \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = carrier G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   179
    and "mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = add G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   180
    and "one \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = zero G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   181
    and "m_inv \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = a_inv G"
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   182
  by (rule a_group) (auto simp: m_inv_def a_inv_def)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   183
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   184
context abelian_group
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   185
begin
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   186
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   187
lemmas a_inv_closed = add.inv_closed
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   188
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   189
lemma minus_closed [intro, simp]:
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   190
  "[| x \<in> carrier G; y \<in> carrier G |] ==> x \<ominus> y \<in> carrier G"
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   191
  by (simp add: a_minus_def)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   192
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   193
lemmas a_l_cancel = add.l_cancel
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   194
lemmas a_r_cancel = add.r_cancel
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   195
lemmas l_neg = add.l_inv [simp del]
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   196
lemmas r_neg = add.r_inv [simp del]
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   197
lemmas minus_zero = add.inv_one
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   198
lemmas minus_minus = add.inv_inv
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   199
lemmas a_inv_inj = add.inv_inj
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   200
lemmas minus_equality = add.inv_equality
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   201
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   202
end
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   203
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   204
sublocale abelian_group <
55926
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   205
  add!: comm_group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   206
  where "carrier \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = carrier G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   207
    and "mult \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = add G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   208
    and "one \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = zero G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   209
    and "m_inv \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = a_inv G"
3ef14caf5637 more symbols;
wenzelm
parents: 48891
diff changeset
   210
    and "finprod \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> = finsum G"
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   211
  by (rule a_comm_group) (auto simp: m_inv_def a_inv_def finsum_def)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   212
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   213
lemmas (in abelian_group) minus_add = add.inv_mult
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   214
 
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   215
text {* Derive an @{text "abelian_group"} from a @{text "comm_group"} *}
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   216
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   217
lemma comm_group_abelian_groupI:
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   218
  fixes G (structure)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   219
  assumes cg: "comm_group \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   220
  shows "abelian_group G"
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   221
proof -
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   222
  interpret comm_group "\<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   223
    by (rule cg)
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   224
  show "abelian_group G" ..
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   225
qed
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   226
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   227
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   228
subsection {* Rings: Basic Definitions *}
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   229
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   230
locale semiring = abelian_monoid R + monoid R for R (structure) +
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   231
  assumes l_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   232
      ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   233
    and r_distr: "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   234
      ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   235
    and l_null[simp]: "x \<in> carrier R ==> \<zero> \<otimes> x = \<zero>"
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   236
    and r_null[simp]: "x \<in> carrier R ==> x \<otimes> \<zero> = \<zero>"
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   237
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   238
locale ring = abelian_group R + monoid R for R (structure) +
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   239
  assumes "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   240
      ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   241
    and "[| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   242
      ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   243
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   244
locale cring = ring + comm_monoid R
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   245
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   246
locale "domain" = cring +
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   247
  assumes one_not_zero [simp]: "\<one> ~= \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   248
    and integral: "[| a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   249
                  a = \<zero> | b = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   250
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   251
locale field = "domain" +
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   252
  assumes field_Units: "Units R = carrier R - {\<zero>}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   253
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   254
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   255
subsection {* Rings *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   256
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   257
lemma ringI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   258
  fixes R (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   259
  assumes abelian_group: "abelian_group R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   260
    and monoid: "monoid R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   261
    and l_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   262
      ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   263
    and r_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   264
      ==> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   265
  shows "ring R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   266
  by (auto intro: ring.intro
27714
27b4d7c01f8b Tuned (for the sake of a meaningless log entry).
ballarin
parents: 27699
diff changeset
   267
    abelian_group.axioms ring_axioms.intro assms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   268
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   269
context ring begin
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   270
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   271
lemma is_abelian_group: "abelian_group R" ..
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   272
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   273
lemma is_monoid: "monoid R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   274
  by (auto intro!: monoidI m_assoc)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   275
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   276
lemma is_ring: "ring R"
26202
51f8a696cd8d explicit referencing of background facts;
wenzelm
parents: 23957
diff changeset
   277
  by (rule ring_axioms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   278
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   279
end
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   280
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   281
lemmas ring_record_simps = monoid_record_simps ring.simps
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   282
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   283
lemma cringI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   284
  fixes R (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   285
  assumes abelian_group: "abelian_group R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   286
    and comm_monoid: "comm_monoid R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   287
    and l_distr: "!!x y z. [| x \<in> carrier R; y \<in> carrier R; z \<in> carrier R |]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   288
      ==> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   289
  shows "cring R"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   290
proof (intro cring.intro ring.intro)
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   291
  show "ring_axioms R"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   292
    -- {* Right-distributivity follows from left-distributivity and
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   293
          commutativity. *}
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   294
  proof (rule ring_axioms.intro)
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   295
    fix x y z
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   296
    assume R: "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R"
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   297
    note [simp] = comm_monoid.axioms [OF comm_monoid]
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   298
      abelian_group.axioms [OF abelian_group]
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   299
      abelian_monoid.a_closed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   300
        
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   301
    from R have "z \<otimes> (x \<oplus> y) = (x \<oplus> y) \<otimes> z"
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   302
      by (simp add: comm_monoid.m_comm [OF comm_monoid.intro])
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   303
    also from R have "... = x \<otimes> z \<oplus> y \<otimes> z" by (simp add: l_distr)
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   304
    also from R have "... = z \<otimes> x \<oplus> z \<otimes> y"
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   305
      by (simp add: comm_monoid.m_comm [OF comm_monoid.intro])
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   306
    finally show "z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" .
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   307
  qed (rule l_distr)
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   308
qed (auto intro: cring.intro
27714
27b4d7c01f8b Tuned (for the sake of a meaningless log entry).
ballarin
parents: 27699
diff changeset
   309
  abelian_group.axioms comm_monoid.axioms ring_axioms.intro assms)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   310
27699
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   311
(*
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   312
lemma (in cring) is_comm_monoid:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   313
  "comm_monoid R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   314
  by (auto intro!: comm_monoidI m_assoc m_comm)
27699
489e3f33af0e New theorems on summation.
ballarin
parents: 27611
diff changeset
   315
*)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   316
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   317
lemma (in cring) is_cring:
26202
51f8a696cd8d explicit referencing of background facts;
wenzelm
parents: 23957
diff changeset
   318
  "cring R" by (rule cring_axioms)
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   319
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   320
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   321
subsubsection {* Normaliser for Rings *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   322
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   323
lemma (in abelian_group) r_neg2:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   324
  "[| x \<in> carrier G; y \<in> carrier G |] ==> x \<oplus> (\<ominus> x \<oplus> y) = y"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   325
proof -
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   326
  assume G: "x \<in> carrier G" "y \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   327
  then have "(x \<oplus> \<ominus> x) \<oplus> y = y"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   328
    by (simp only: r_neg l_zero)
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   329
  with G show ?thesis
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   330
    by (simp add: a_ac)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   331
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   332
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   333
lemma (in abelian_group) r_neg1:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   334
  "[| x \<in> carrier G; y \<in> carrier G |] ==> \<ominus> x \<oplus> (x \<oplus> y) = y"
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 G: "x \<in> carrier G" "y \<in> carrier G"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   337
  then have "(\<ominus> x \<oplus> x) \<oplus> y = y" 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   338
    by (simp only: l_neg l_zero)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   339
  with G show ?thesis by (simp add: a_ac)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   340
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   341
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   342
context ring begin
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   343
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   344
text {* 
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   345
  The following proofs are from Jacobson, Basic Algebra I, pp.~88--89.
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   346
*}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   347
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   348
sublocale semiring
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   349
proof -
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   350
  note [simp] = ring_axioms[unfolded ring_def ring_axioms_def]
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   351
  show "semiring R"
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   352
  proof (unfold_locales)
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   353
    fix x
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   354
    assume R: "x \<in> carrier R"
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   355
    then have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = (\<zero> \<oplus> \<zero>) \<otimes> x"
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   356
      by (simp del: l_zero r_zero)
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   357
    also from R have "... = \<zero> \<otimes> x \<oplus> \<zero>" by simp
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   358
    finally have "\<zero> \<otimes> x \<oplus> \<zero> \<otimes> x = \<zero> \<otimes> x \<oplus> \<zero>" .
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   359
    with R show "\<zero> \<otimes> x = \<zero>" by (simp del: r_zero)
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   360
    from R have "x \<otimes> \<zero> \<oplus> x \<otimes> \<zero> = x \<otimes> (\<zero> \<oplus> \<zero>)"
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   361
      by (simp del: l_zero r_zero)
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   362
    also from R have "... = x \<otimes> \<zero> \<oplus> \<zero>" by simp
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   363
    finally have "x \<otimes> \<zero> \<oplus> x \<otimes> \<zero> = x \<otimes> \<zero> \<oplus> \<zero>" .
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   364
    with R show "x \<otimes> \<zero> = \<zero>" by (simp del: r_zero)
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   365
  qed auto
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   366
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   367
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   368
lemma l_minus:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   369
  "[| x \<in> carrier R; y \<in> carrier R |] ==> \<ominus> x \<otimes> y = \<ominus> (x \<otimes> y)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   370
proof -
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   371
  assume R: "x \<in> carrier R" "y \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   372
  then have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = (\<ominus> x \<oplus> x) \<otimes> y" by (simp add: l_distr)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 41959
diff changeset
   373
  also from R have "... = \<zero>" by (simp add: l_neg)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   374
  finally have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = \<zero>" .
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   375
  with R have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp
21896
9a7949815a84 Experimenting with interpretations of "definition".
ballarin
parents: 20318
diff changeset
   376
  with R show ?thesis by (simp add: a_assoc r_neg)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   377
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   378
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   379
lemma r_minus:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   380
  "[| x \<in> carrier R; y \<in> carrier R |] ==> x \<otimes> \<ominus> y = \<ominus> (x \<otimes> y)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   381
proof -
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   382
  assume R: "x \<in> carrier R" "y \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   383
  then have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = x \<otimes> (\<ominus> y \<oplus> y)" by (simp add: r_distr)
44677
3fb27b19e058 misc tuning and simplification of proofs;
wenzelm
parents: 41959
diff changeset
   384
  also from R have "... = \<zero>" by (simp add: l_neg)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   385
  finally have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = \<zero>" .
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   386
  with R have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y \<oplus> \<ominus> (x \<otimes> y) = \<zero> \<oplus> \<ominus> (x \<otimes> y)" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   387
  with R show ?thesis by (simp add: a_assoc r_neg )
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   388
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   389
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   390
end
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   391
23957
54fab60ddc97 Interpretation of rings (as integers) maps defined operations to defined
ballarin
parents: 23350
diff changeset
   392
lemma (in abelian_group) minus_eq:
54fab60ddc97 Interpretation of rings (as integers) maps defined operations to defined
ballarin
parents: 23350
diff changeset
   393
  "[| x \<in> carrier G; y \<in> carrier G |] ==> x \<ominus> y = x \<oplus> \<ominus> y"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   394
  by (simp only: a_minus_def)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   395
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   396
text {* Setup algebra method:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   397
  compute distributive normal form in locale contexts *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   398
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47701
diff changeset
   399
ML_file "ringsimp.ML"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   400
58811
19382bbfa93a modernized setup;
wenzelm
parents: 55926
diff changeset
   401
attribute_setup algebra = {*
19382bbfa93a modernized setup;
wenzelm
parents: 55926
diff changeset
   402
  Scan.lift ((Args.add >> K true || Args.del >> K false) --| Args.colon || Scan.succeed true)
19382bbfa93a modernized setup;
wenzelm
parents: 55926
diff changeset
   403
    -- Scan.lift Args.name -- Scan.repeat Args.term
19382bbfa93a modernized setup;
wenzelm
parents: 55926
diff changeset
   404
    >> (fn ((b, n), ts) => if b then Ringsimp.add_struct (n, ts) else Ringsimp.del_struct (n, ts))
19382bbfa93a modernized setup;
wenzelm
parents: 55926
diff changeset
   405
*} "theorems controlling algebra method"
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47409
diff changeset
   406
157e6108a342 more standard method setup;
wenzelm
parents: 47409
diff changeset
   407
method_setup algebra = {*
58811
19382bbfa93a modernized setup;
wenzelm
parents: 55926
diff changeset
   408
  Scan.succeed (SIMPLE_METHOD' o Ringsimp.algebra_tac)
47701
157e6108a342 more standard method setup;
wenzelm
parents: 47409
diff changeset
   409
*} "normalisation of algebraic structure"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   410
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   411
lemmas (in semiring) semiring_simprules
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   412
  [algebra ring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] =
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   413
  a_closed zero_closed  m_closed one_closed
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   414
  a_assoc l_zero  a_comm m_assoc l_one l_distr r_zero
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   415
  a_lcomm r_distr l_null r_null 
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   416
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   417
lemmas (in ring) ring_simprules
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   418
  [algebra ring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] =
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   419
  a_closed zero_closed a_inv_closed minus_closed m_closed one_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   420
  a_assoc l_zero l_neg a_comm m_assoc l_one l_distr minus_eq
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   421
  r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   422
  a_lcomm r_distr l_null r_null l_minus r_minus
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   423
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   424
lemmas (in cring)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   425
  [algebra del: ring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] =
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   426
  _
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   427
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   428
lemmas (in cring) cring_simprules
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   429
  [algebra add: cring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] =
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   430
  a_closed zero_closed a_inv_closed minus_closed m_closed one_closed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   431
  a_assoc l_zero l_neg a_comm m_assoc l_one l_distr m_comm minus_eq
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   432
  r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   433
  a_lcomm m_lcomm r_distr l_null r_null l_minus r_minus
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   434
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   435
lemma (in semiring) nat_pow_zero:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   436
  "(n::nat) ~= 0 ==> \<zero> (^) n = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   437
  by (induct n) simp_all
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   438
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   439
context semiring begin
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   440
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   441
lemma one_zeroD:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   442
  assumes onezero: "\<one> = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   443
  shows "carrier R = {\<zero>}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   444
proof (rule, rule)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   445
  fix x
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   446
  assume xcarr: "x \<in> carrier R"
47409
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   447
  from xcarr have "x = x \<otimes> \<one>" by simp
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   448
  with onezero have "x = x \<otimes> \<zero>" by simp
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   449
  with xcarr have "x = \<zero>" by simp
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   450
  then show "x \<in> {\<zero>}" by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   451
qed fast
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   452
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   453
lemma one_zeroI:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   454
  assumes carrzero: "carrier R = {\<zero>}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   455
  shows "\<one> = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   456
proof -
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   457
  from one_closed and carrzero
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   458
      show "\<one> = \<zero>" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   459
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   460
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   461
lemma carrier_one_zero: "(carrier R = {\<zero>}) = (\<one> = \<zero>)"
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   462
  apply rule
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   463
   apply (erule one_zeroI)
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   464
  apply (erule one_zeroD)
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   465
  done
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   466
46721
f88b187ad8ca tuned proofs;
wenzelm
parents: 44677
diff changeset
   467
lemma carrier_one_not_zero: "(carrier R \<noteq> {\<zero>}) = (\<one> \<noteq> \<zero>)"
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   468
  by (simp add: carrier_one_zero)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   469
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   470
end
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   471
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   472
text {* Two examples for use of method algebra *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   473
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   474
lemma
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   475
  fixes R (structure) and S (structure)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   476
  assumes "ring R" "cring S"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   477
  assumes RS: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier S" "d \<in> carrier S"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   478
  shows "a \<oplus> \<ominus> (a \<oplus> \<ominus> b) = b & c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   479
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
   480
  interpret ring R by fact
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
   481
  interpret cring S by fact
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   482
  from RS show ?thesis by algebra
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   483
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   484
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   485
lemma
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   486
  fixes R (structure)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   487
  assumes "ring R"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   488
  assumes R: "a \<in> carrier R" "b \<in> carrier R"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   489
  shows "a \<ominus> (a \<ominus> b) = b"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   490
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
   491
  interpret ring R by fact
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   492
  from R show ?thesis by algebra
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26202
diff changeset
   493
qed
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   494
35849
b5522b51cb1e standard headers;
wenzelm
parents: 35848
diff changeset
   495
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   496
subsubsection {* Sums over Finite Sets *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   497
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   498
lemma (in semiring) finsum_ldistr:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   499
  "[| finite A; a \<in> carrier R; f \<in> A -> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   500
   finsum R f A \<otimes> a = finsum R (%i. f i \<otimes> a) A"
22265
3c5c6bdf61de Adapted to changes in Finite_Set theory.
berghofe
parents: 21896
diff changeset
   501
proof (induct set: finite)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   502
  case empty then show ?case by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   503
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   504
  case (insert x F) then show ?case by (simp add: Pi_def l_distr)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   505
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   506
59851
43b1870b3e61 added locale for semirings
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 58811
diff changeset
   507
lemma (in semiring) finsum_rdistr:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   508
  "[| finite A; a \<in> carrier R; f \<in> A -> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   509
   a \<otimes> finsum R f A = finsum R (%i. a \<otimes> f i) A"
22265
3c5c6bdf61de Adapted to changes in Finite_Set theory.
berghofe
parents: 21896
diff changeset
   510
proof (induct set: finite)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   511
  case empty then show ?case by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   512
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   513
  case (insert x F) then show ?case by (simp add: Pi_def r_distr)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   514
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   515
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   516
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   517
subsection {* Integral Domains *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   518
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   519
context "domain" begin
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   520
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   521
lemma zero_not_one [simp]:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   522
  "\<zero> ~= \<one>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   523
  by (rule not_sym) simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   524
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   525
lemma integral_iff: (* not by default a simp rule! *)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   526
  "[| a \<in> carrier R; b \<in> carrier R |] ==> (a \<otimes> b = \<zero>) = (a = \<zero> | b = \<zero>)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   527
proof
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   528
  assume "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   529
  then show "a = \<zero> | b = \<zero>" by (simp add: integral)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   530
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   531
  assume "a \<in> carrier R" "b \<in> carrier R" "a = \<zero> | b = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   532
  then show "a \<otimes> b = \<zero>" by auto
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   533
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   534
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   535
lemma m_lcancel:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   536
  assumes prem: "a ~= \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   537
    and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   538
  shows "(a \<otimes> b = a \<otimes> c) = (b = c)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   539
proof
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   540
  assume eq: "a \<otimes> b = a \<otimes> c"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   541
  with R have "a \<otimes> (b \<ominus> c) = \<zero>" by algebra
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   542
  with R have "a = \<zero> | (b \<ominus> c) = \<zero>" by (simp add: integral_iff)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   543
  with prem and R have "b \<ominus> c = \<zero>" by auto 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   544
  with R have "b = b \<ominus> (b \<ominus> c)" by algebra 
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   545
  also from R have "b \<ominus> (b \<ominus> c) = c" by algebra
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   546
  finally show "b = c" .
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   547
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   548
  assume "b = c" then show "a \<otimes> b = a \<otimes> c" by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   549
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   550
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   551
lemma m_rcancel:
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   552
  assumes prem: "a ~= \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   553
    and R: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   554
  shows conc: "(b \<otimes> a = c \<otimes> a) = (b = c)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   555
proof -
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   556
  from prem and R have "(a \<otimes> b = a \<otimes> c) = (b = c)" by (rule m_lcancel)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   557
  with R show ?thesis by algebra
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   558
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   559
41433
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   560
end
1b8ff770f02c Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents: 35849
diff changeset
   561
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   562
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   563
subsection {* Fields *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   564
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   565
text {* Field would not need to be derived from domain, the properties
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   566
  for domain follow from the assumptions of field *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   567
lemma (in cring) cring_fieldI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   568
  assumes field_Units: "Units R = carrier R - {\<zero>}"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   569
  shows "field R"
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27933
diff changeset
   570
proof
47409
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   571
  from field_Units have "\<zero> \<notin> Units R" by fast
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   572
  moreover have "\<one> \<in> Units R" by fast
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   573
  ultimately show "\<one> \<noteq> \<zero>" by force
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   574
next
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   575
  fix a b
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   576
  assume acarr: "a \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   577
    and bcarr: "b \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   578
    and ab: "a \<otimes> b = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   579
  show "a = \<zero> \<or> b = \<zero>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   580
  proof (cases "a = \<zero>", simp)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   581
    assume "a \<noteq> \<zero>"
47409
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   582
    with field_Units and acarr have aUnit: "a \<in> Units R" by fast
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   583
    from bcarr have "b = \<one> \<otimes> b" by algebra
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   584
    also from aUnit acarr have "... = (inv a \<otimes> a) \<otimes> b" by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   585
    also from acarr bcarr aUnit[THEN Units_inv_closed]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   586
    have "... = (inv a) \<otimes> (a \<otimes> b)" by algebra
47409
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   587
    also from ab and acarr bcarr aUnit have "... = (inv a) \<otimes> \<zero>" by simp
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   588
    also from aUnit[THEN Units_inv_closed] have "... = \<zero>" by algebra
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   589
    finally have "b = \<zero>" .
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   590
    then show "a = \<zero> \<or> b = \<zero>" by simp
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   591
  qed
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22265
diff changeset
   592
qed (rule field_Units)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   593
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   594
text {* Another variant to show that something is a field *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   595
lemma (in cring) cring_fieldI2:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   596
  assumes notzero: "\<zero> \<noteq> \<one>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   597
  and invex: "\<And>a. \<lbrakk>a \<in> carrier R; a \<noteq> \<zero>\<rbrakk> \<Longrightarrow> \<exists>b\<in>carrier R. a \<otimes> b = \<one>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   598
  shows "field R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   599
  apply (rule cring_fieldI, simp add: Units_def)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   600
  apply (rule, clarsimp)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   601
  apply (simp add: notzero)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   602
proof (clarsimp)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   603
  fix x
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   604
  assume xcarr: "x \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   605
    and "x \<noteq> \<zero>"
47409
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   606
  then have "\<exists>y\<in>carrier R. x \<otimes> y = \<one>" by (rule invex)
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   607
  then obtain y where ycarr: "y \<in> carrier R" and xy: "x \<otimes> y = \<one>" by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   608
  from xy xcarr ycarr have "y \<otimes> x = \<one>" by (simp add: m_comm)
47409
c5be1120980d tuned proofs;
wenzelm
parents: 46721
diff changeset
   609
  with ycarr and xy show "\<exists>y\<in>carrier R. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>" by fast
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   610
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   611
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   612
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   613
subsection {* Morphisms *}
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   614
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
   615
definition
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   616
  ring_hom :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme] => ('a => 'b) set"
35848
5443079512ea slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents: 35847
diff changeset
   617
  where "ring_hom R S =
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
   618
    {h. h \<in> carrier R -> carrier S &
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   619
      (ALL x y. x \<in> carrier R & y \<in> carrier R -->
35847
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
   620
        h (x \<otimes>\<^bsub>R\<^esub> y) = h x \<otimes>\<^bsub>S\<^esub> h y & h (x \<oplus>\<^bsub>R\<^esub> y) = h x \<oplus>\<^bsub>S\<^esub> h y) &
19f1f7066917 eliminated old constdefs;
wenzelm
parents: 35416
diff changeset
   621
      h \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>}"
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   622
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   623
lemma ring_hom_memI:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   624
  fixes R (structure) and S (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   625
  assumes hom_closed: "!!x. x \<in> carrier R ==> h x \<in> carrier S"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   626
    and hom_mult: "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   627
      h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   628
    and hom_add: "!!x y. [| x \<in> carrier R; y \<in> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   629
      h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   630
    and hom_one: "h \<one> = \<one>\<^bsub>S\<^esub>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   631
  shows "h \<in> ring_hom R S"
27714
27b4d7c01f8b Tuned (for the sake of a meaningless log entry).
ballarin
parents: 27699
diff changeset
   632
  by (auto simp add: ring_hom_def assms Pi_def)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   633
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   634
lemma ring_hom_closed:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   635
  "[| h \<in> ring_hom R S; x \<in> carrier R |] ==> h x \<in> carrier S"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   636
  by (auto simp add: ring_hom_def funcset_mem)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   637
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   638
lemma ring_hom_mult:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   639
  fixes R (structure) and S (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   640
  shows
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   641
    "[| h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   642
    h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   643
    by (simp add: ring_hom_def)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   644
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   645
lemma ring_hom_add:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   646
  fixes R (structure) and S (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   647
  shows
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   648
    "[| h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R |] ==>
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   649
    h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   650
    by (simp add: ring_hom_def)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   651
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   652
lemma ring_hom_one:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   653
  fixes R (structure) and S (structure)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   654
  shows "h \<in> ring_hom R S ==> h \<one> = \<one>\<^bsub>S\<^esub>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   655
  by (simp add: ring_hom_def)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   656
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
   657
locale ring_hom_cring = R: cring R + S: cring S
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
   658
    for R (structure) and S (structure) +
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   659
  fixes h
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   660
  assumes homh [simp, intro]: "h \<in> ring_hom R S"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   661
  notes hom_closed [simp, intro] = ring_hom_closed [OF homh]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   662
    and hom_mult [simp] = ring_hom_mult [OF homh]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   663
    and hom_add [simp] = ring_hom_add [OF homh]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   664
    and hom_one [simp] = ring_hom_one [OF homh]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   665
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   666
lemma (in ring_hom_cring) hom_zero [simp]:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   667
  "h \<zero> = \<zero>\<^bsub>S\<^esub>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   668
proof -
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   669
  have "h \<zero> \<oplus>\<^bsub>S\<^esub> h \<zero> = h \<zero> \<oplus>\<^bsub>S\<^esub> \<zero>\<^bsub>S\<^esub>"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   670
    by (simp add: hom_add [symmetric] del: hom_add)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   671
  then show ?thesis by (simp del: S.r_zero)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   672
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   673
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   674
lemma (in ring_hom_cring) hom_a_inv [simp]:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   675
  "x \<in> carrier R ==> h (\<ominus> x) = \<ominus>\<^bsub>S\<^esub> h x"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   676
proof -
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   677
  assume R: "x \<in> carrier R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   678
  then have "h x \<oplus>\<^bsub>S\<^esub> h (\<ominus> x) = h x \<oplus>\<^bsub>S\<^esub> (\<ominus>\<^bsub>S\<^esub> h x)"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   679
    by (simp add: hom_add [symmetric] R.r_neg S.r_neg del: hom_add)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   680
  with R show ?thesis by simp
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   681
qed
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   682
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   683
lemma (in ring_hom_cring) hom_finsum [simp]:
60112
3eab4acaa035 finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 59851
diff changeset
   684
  "f \<in> A -> carrier R ==>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   685
  h (finsum R f A) = finsum S (h o f) A"
60112
3eab4acaa035 finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 59851
diff changeset
   686
  by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   687
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   688
lemma (in ring_hom_cring) hom_finprod:
60112
3eab4acaa035 finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 59851
diff changeset
   689
  "f \<in> A -> carrier R ==>
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   690
  h (finprod R f A) = finprod S (h o f) A"
60112
3eab4acaa035 finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents: 59851
diff changeset
   691
  by (induct A rule: infinite_finite_induct, auto simp: Pi_def)
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   692
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   693
declare ring_hom_cring.hom_finprod [simp]
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   694
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   695
lemma id_ring_hom [simp]:
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   696
  "id \<in> ring_hom R R"
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   697
  by (auto intro!: ring_hom_memI)
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   698
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
diff changeset
   699
end