| author | nipkow | 
| Sat, 14 Jun 2025 11:45:56 +0200 | |
| changeset 82704 | e0fb46018187 | 
| parent 81142 | 6ad2c917dd2e | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Algebra/Ring.thy | 
| 35849 | 2 | Author: Clemens Ballarin, started 9 December 1996 | 
| 68552 | 3 | |
| 68582 | 4 | With contributions by Martin Baillon. | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 5 | *) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 6 | |
| 28823 | 7 | theory Ring | 
| 8 | imports FiniteProduct | |
| 35847 | 9 | begin | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 10 | |
| 61382 | 11 | section \<open>The Algebraic Hierarchy of Rings\<close> | 
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 12 | |
| 61382 | 13 | subsection \<open>Abelian Groups\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 14 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 15 | record 'a ring = "'a monoid" + | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 16 | zero :: 'a (\<open>\<zero>\<index>\<close>) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
80768diff
changeset | 17 | add :: "['a, 'a] \<Rightarrow> 'a" (infixl \<open>\<oplus>\<index>\<close> 65) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 18 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 19 | abbreviation | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 20 |   add_monoid :: "('a, 'm) ring_scheme \<Rightarrow> ('a, 'm) monoid_scheme"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 21 | where "add_monoid R \<equiv> \<lparr> carrier = carrier R, mult = add R, one = zero R, \<dots> = (undefined :: 'm) \<rparr>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 22 | |
| 61382 | 23 | text \<open>Derived operations.\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 24 | |
| 35847 | 25 | definition | 
| 81142 | 26 |   a_inv :: "[('a, 'm) ring_scheme, 'a ] \<Rightarrow> 'a" (\<open>(\<open>open_block notation=\<open>prefix \<ominus>\<close>\<close>\<ominus>\<index> _)\<close> [81] 80)
 | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 27 | where "a_inv R = m_inv (add_monoid R)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 28 | |
| 35847 | 29 | definition | 
| 81142 | 30 |   a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" (\<open>(\<open>notation=\<open>infix \<ominus>\<close>\<close>_ \<ominus>\<index> _)\<close> [65,66] 65)
 | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 31 | where "x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 32 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 33 | definition | 
| 81142 | 34 |   add_pow :: "[_, ('b :: semiring_1), 'a] \<Rightarrow> 'a"
 | 
| 35 | (\<open>(\<open>open_block notation=\<open>mixfix \<cdot>\<close>\<close>[_] \<cdot>\<index> _)\<close> [81, 81] 80) | |
| 36 | where "[k] \<cdot>\<^bsub>R\<^esub> a = pow (add_monoid R) a k" | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 37 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 38 | locale abelian_monoid = | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 39 | fixes G (structure) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 40 | assumes a_comm_monoid: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 41 | "comm_monoid (add_monoid G)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 42 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 43 | definition | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 44 |   finsum :: "[('b, 'm) ring_scheme, 'a \<Rightarrow> 'b, 'a set] \<Rightarrow> 'b" where
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 45 | "finsum G = finprod (add_monoid G)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 46 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 47 | syntax | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 48 | "_finsum" :: "index \<Rightarrow> idt \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" | 
| 81142 | 49 | (\<open>(\<open>indent=3 notation=\<open>binder \<Oplus>\<close>\<close>\<Oplus>__\<in>_. _)\<close> [1000, 0, 51, 10] 10) | 
| 80768 | 50 | syntax_consts | 
| 51 | "_finsum" \<rightleftharpoons> finsum | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 52 | translations | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 53 | "\<Oplus>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finsum G (\<lambda>i. b) A" | 
| 63167 | 54 | \<comment> \<open>Beware of argument permutation!\<close> | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 55 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 56 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 57 | locale abelian_group = abelian_monoid + | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 58 | assumes a_comm_group: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 59 | "comm_group (add_monoid G)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 60 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 61 | |
| 61382 | 62 | subsection \<open>Basic Properties\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 63 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 64 | lemma abelian_monoidI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 65 | fixes R (structure) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 66 | assumes "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 67 | and "\<zero> \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 68 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 69 | and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 70 | and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 71 | shows "abelian_monoid R" | 
| 27714 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 ballarin parents: 
27699diff
changeset | 72 | by (auto intro!: abelian_monoid.intro comm_monoidI intro: assms) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 73 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 74 | lemma abelian_monoidE: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 75 | fixes R (structure) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 76 | assumes "abelian_monoid R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 77 | shows "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 78 | and "\<zero> \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 79 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 80 | and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 81 | and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 82 | using assms unfolding abelian_monoid_def comm_monoid_def comm_monoid_axioms_def monoid_def by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 83 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 84 | lemma abelian_groupI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 85 | fixes R (structure) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 86 | assumes "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 87 | and "\<zero> \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 88 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 89 | and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 90 | and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 91 | and "\<And>x. x \<in> carrier R \<Longrightarrow> \<exists>y \<in> carrier R. y \<oplus> x = \<zero>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 92 | shows "abelian_group R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 93 | by (auto intro!: abelian_group.intro abelian_monoidI | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 94 | abelian_group_axioms.intro comm_monoidI comm_groupI | 
| 27714 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 ballarin parents: 
27699diff
changeset | 95 | intro: assms) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 96 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 97 | lemma abelian_groupE: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 98 | fixes R (structure) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 99 | assumes "abelian_group R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 100 | shows "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 101 | and "\<zero> \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 102 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 103 | and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<oplus> y = y \<oplus> x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 104 | and "\<And>x. x \<in> carrier R \<Longrightarrow> \<zero> \<oplus> x = x" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 105 | and "\<And>x. x \<in> carrier R \<Longrightarrow> \<exists>y \<in> carrier R. y \<oplus> x = \<zero>" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 106 | using abelian_group.a_comm_group assms comm_groupE by fastforce+ | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 107 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 108 | lemma (in abelian_monoid) a_monoid: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 109 | "monoid (add_monoid G)" | 
| 68517 | 110 | by (rule comm_monoid.axioms, rule a_comm_monoid) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 111 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 112 | lemma (in abelian_group) a_group: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 113 | "group (add_monoid G)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 114 | by (simp add: group_def a_monoid) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 115 | (simp add: comm_group.axioms group.axioms a_comm_group) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 116 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 117 | lemmas monoid_record_simps = partial_object.simps monoid.simps | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 118 | |
| 61382 | 119 | text \<open>Transfer facts from multiplicative structures via interpretation.\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 120 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 121 | sublocale abelian_monoid < | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 122 | add: monoid "(add_monoid G)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 123 | rewrites "carrier (add_monoid G) = carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 124 | and "mult (add_monoid G) = add G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 125 | and "one (add_monoid G) = zero G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 126 | and "(\<lambda>a k. pow (add_monoid G) a k) = (\<lambda>a k. add_pow G k a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 127 | by (rule a_monoid) (auto simp add: add_pow_def) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 128 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 129 | context abelian_monoid | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 130 | begin | 
| 27933 | 131 | |
| 68517 | 132 | lemmas a_closed = add.m_closed | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 133 | lemmas zero_closed = add.one_closed | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 134 | lemmas a_assoc = add.m_assoc | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 135 | lemmas l_zero = add.l_one | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 136 | lemmas r_zero = add.r_one | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 137 | lemmas minus_unique = add.inv_unique | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 138 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 139 | end | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 140 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 141 | sublocale abelian_monoid < | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 142 | add: comm_monoid "(add_monoid G)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 143 | rewrites "carrier (add_monoid G) = carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 144 | and "mult (add_monoid G) = add G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 145 | and "one (add_monoid G) = zero G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 146 | and "finprod (add_monoid G) = finsum G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 147 | and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 148 | by (rule a_comm_monoid) (auto simp: finsum_def add_pow_def) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 149 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 150 | context abelian_monoid begin | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 151 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 152 | lemmas a_comm = add.m_comm | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 153 | lemmas a_lcomm = add.m_lcomm | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 154 | lemmas a_ac = a_assoc a_comm a_lcomm | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 155 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 156 | lemmas finsum_empty = add.finprod_empty | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 157 | lemmas finsum_insert = add.finprod_insert | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 158 | lemmas finsum_zero = add.finprod_one | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 159 | lemmas finsum_closed = add.finprod_closed | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 160 | lemmas finsum_Un_Int = add.finprod_Un_Int | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 161 | lemmas finsum_Un_disjoint = add.finprod_Un_disjoint | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 162 | lemmas finsum_addf = add.finprod_multf | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 163 | lemmas finsum_cong' = add.finprod_cong' | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 164 | lemmas finsum_0 = add.finprod_0 | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 165 | lemmas finsum_Suc = add.finprod_Suc | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 166 | lemmas finsum_Suc2 = add.finprod_Suc2 | 
| 60112 
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
59851diff
changeset | 167 | lemmas finsum_infinite = add.finprod_infinite | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 168 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 169 | lemmas finsum_cong = add.finprod_cong | 
| 61382 | 170 | text \<open>Usually, if this rule causes a failed congruence proof error, | 
| 63167 | 171 | the reason is that the premise \<open>g \<in> B \<rightarrow> carrier G\<close> cannot be shown. | 
| 61382 | 172 |    Adding @{thm [source] Pi_def} to the simpset is often useful.\<close>
 | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 173 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 174 | lemmas finsum_reindex = add.finprod_reindex | 
| 27699 | 175 | |
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67091diff
changeset | 176 | (* The following would be wrong. Needed is the equivalent of [^] for addition, | 
| 27699 | 177 | or indeed the canonical embedding from Nat into the monoid. | 
| 178 | ||
| 27933 | 179 | lemma finsum_const: | 
| 27699 | 180 | assumes fin [simp]: "finite A" | 
| 181 | and a [simp]: "a : carrier G" | |
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67091diff
changeset | 182 | shows "finsum G (%x. a) A = a [^] card A" | 
| 27699 | 183 | using fin apply induct | 
| 184 | apply force | |
| 185 | apply (subst finsum_insert) | |
| 186 | apply auto | |
| 187 | apply (force simp add: Pi_def) | |
| 188 | apply (subst m_comm) | |
| 189 | apply auto | |
| 190 | done | |
| 191 | *) | |
| 192 | ||
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 193 | lemmas finsum_singleton = add.finprod_singleton | 
| 27933 | 194 | |
| 195 | end | |
| 196 | ||
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 197 | sublocale abelian_group < | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 198 | add: group "(add_monoid G)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 199 | rewrites "carrier (add_monoid G) = carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 200 | and "mult (add_monoid G) = add G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 201 | and "one (add_monoid G) = zero G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 202 | and "m_inv (add_monoid G) = a_inv G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 203 | and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 204 | by (rule a_group) (auto simp: m_inv_def a_inv_def add_pow_def) | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 205 | |
| 55926 | 206 | context abelian_group | 
| 207 | begin | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 208 | |
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 209 | lemmas a_inv_closed = add.inv_closed | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 210 | |
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 211 | lemma minus_closed [intro, simp]: | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 212 | "[| 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: 
35849diff
changeset | 213 | by (simp add: a_minus_def) | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 214 | |
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 215 | lemmas l_neg = add.l_inv [simp del] | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 216 | lemmas r_neg = add.r_inv [simp del] | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 217 | lemmas minus_minus = add.inv_inv | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 218 | lemmas a_inv_inj = add.inv_inj | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 219 | lemmas minus_equality = add.inv_equality | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 220 | |
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 221 | end | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 222 | |
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 223 | sublocale abelian_group < | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 224 | add: comm_group "(add_monoid G)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 225 | rewrites "carrier (add_monoid G) = carrier G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 226 | and "mult (add_monoid G) = add G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 227 | and "one (add_monoid G) = zero G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 228 | and "m_inv (add_monoid G) = a_inv G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 229 | and "finprod (add_monoid G) = finsum G" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 230 | and "pow (add_monoid G) = (\<lambda>a k. add_pow G k a)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 231 | by (rule a_comm_group) (auto simp: m_inv_def a_inv_def finsum_def add_pow_def) | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 232 | |
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 233 | lemmas (in abelian_group) minus_add = add.inv_mult | 
| 68517 | 234 | |
| 63167 | 235 | text \<open>Derive an \<open>abelian_group\<close> from a \<open>comm_group\<close>\<close> | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 236 | |
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 237 | lemma comm_group_abelian_groupI: | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 238 | fixes G (structure) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 239 | assumes cg: "comm_group (add_monoid G)" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 240 | shows "abelian_group G" | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 241 | proof - | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 242 | interpret comm_group "(add_monoid G)" | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 243 | by (rule cg) | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 244 | show "abelian_group G" .. | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 245 | qed | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 246 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 247 | |
| 61382 | 248 | subsection \<open>Rings: Basic Definitions\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 249 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 250 | locale semiring = abelian_monoid (* for add *) R + monoid (* for mult *) R for R (structure) + | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 251 | assumes l_distr: "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 252 | and r_distr: "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 253 | and l_null[simp]: "x \<in> carrier R \<Longrightarrow> \<zero> \<otimes> x = \<zero>" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 254 | and r_null[simp]: "x \<in> carrier R \<Longrightarrow> x \<otimes> \<zero> = \<zero>" | 
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 255 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 256 | locale ring = abelian_group (* for add *) R + monoid (* for mult *) R for R (structure) + | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 257 | assumes "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 258 | and "\<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> 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 | 259 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 260 | locale cring = ring + comm_monoid (* for mult *) R | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 261 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 262 | locale "domain" = cring + | 
| 67091 | 263 | assumes one_not_zero [simp]: "\<one> \<noteq> \<zero>" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 264 | and integral: "\<lbrakk> a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> a = \<zero> \<or> b = \<zero>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 265 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 266 | locale field = "domain" + | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 267 |   assumes field_Units: "Units R = carrier R - {\<zero>}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 268 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 269 | |
| 61382 | 270 | subsection \<open>Rings\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 271 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 272 | lemma ringI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 273 | fixes R (structure) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 274 | assumes "abelian_group R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 275 | and "monoid R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 276 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 277 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> 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 | 278 | shows "ring R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 279 | by (auto intro: ring.intro | 
| 27714 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 ballarin parents: 
27699diff
changeset | 280 | abelian_group.axioms ring_axioms.intro assms) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 281 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 282 | lemma ringE: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 283 | fixes R (structure) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 284 | assumes "ring R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 285 | shows "abelian_group R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 286 | and "monoid R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 287 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 288 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 289 | using assms unfolding ring_def ring_axioms_def by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 290 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 291 | context ring begin | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 292 | |
| 46721 | 293 | lemma is_abelian_group: "abelian_group R" .. | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 294 | |
| 46721 | 295 | lemma is_monoid: "monoid R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 296 | by (auto intro!: monoidI m_assoc) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 297 | |
| 75963 
884dbbc8e1b3
avoid duplicate fact error on global_interpretation of residues
 haftmann parents: 
70214diff
changeset | 298 | end | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 299 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 300 | thm monoid_record_simps | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 301 | lemmas ring_record_simps = monoid_record_simps ring.simps | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 302 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 303 | lemma cringI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 304 | fixes R (structure) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 305 | assumes abelian_group: "abelian_group R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 306 | and comm_monoid: "comm_monoid R" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 307 | and l_distr: "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 308 | (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 309 | shows "cring R" | 
| 23350 | 310 | proof (intro cring.intro ring.intro) | 
| 311 | show "ring_axioms R" | |
| 63167 | 312 | \<comment> \<open>Right-distributivity follows from left-distributivity and | 
| 61382 | 313 | commutativity.\<close> | 
| 23350 | 314 | proof (rule ring_axioms.intro) | 
| 315 | fix x y z | |
| 316 | assume R: "x \<in> carrier R" "y \<in> carrier R" "z \<in> carrier R" | |
| 317 | note [simp] = comm_monoid.axioms [OF comm_monoid] | |
| 318 | abelian_group.axioms [OF abelian_group] | |
| 319 | abelian_monoid.a_closed | |
| 68517 | 320 | |
| 23350 | 321 | from R have "z \<otimes> (x \<oplus> y) = (x \<oplus> y) \<otimes> z" | 
| 322 | by (simp add: comm_monoid.m_comm [OF comm_monoid.intro]) | |
| 323 | also from R have "... = x \<otimes> z \<oplus> y \<otimes> z" by (simp add: l_distr) | |
| 324 | also from R have "... = z \<otimes> x \<oplus> z \<otimes> y" | |
| 325 | by (simp add: comm_monoid.m_comm [OF comm_monoid.intro]) | |
| 326 | finally show "z \<otimes> (x \<oplus> y) = z \<otimes> x \<oplus> z \<otimes> y" . | |
| 327 | qed (rule l_distr) | |
| 328 | qed (auto intro: cring.intro | |
| 27714 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 ballarin parents: 
27699diff
changeset | 329 | abelian_group.axioms comm_monoid.axioms ring_axioms.intro assms) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 330 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 331 | lemma cringE: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 332 | fixes R (structure) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 333 | assumes "cring R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 334 | shows "comm_monoid R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 335 | and "\<And>x y z. \<lbrakk> x \<in> carrier R; y \<in> carrier R; z \<in> carrier R \<rbrakk> \<Longrightarrow> (x \<oplus> y) \<otimes> z = x \<otimes> z \<oplus> y \<otimes> z" | 
| 68673 | 336 | using assms cring_def by auto (simp add: assms cring.axioms(1) ringE(3)) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 337 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 338 | lemma (in cring) is_cring: | 
| 26202 | 339 | "cring R" by (rule cring_axioms) | 
| 23350 | 340 | |
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 341 | lemma (in ring) minus_zero [simp]: "\<ominus> \<zero> = \<zero>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 342 | by (simp add: a_inv_def) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 343 | |
| 61382 | 344 | subsubsection \<open>Normaliser for Rings\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 345 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 346 | lemma (in abelian_group) r_neg1: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 347 | "\<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> (\<ominus> x) \<oplus> (x \<oplus> y) = y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 348 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 349 | assume G: "x \<in> carrier G" "y \<in> carrier G" | 
| 68517 | 350 | then have "(\<ominus> x \<oplus> x) \<oplus> y = y" | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 351 | by (simp only: l_neg l_zero) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 352 | with G show ?thesis by (simp add: a_ac) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 353 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 354 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 355 | lemma (in abelian_group) r_neg2: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 356 | "\<lbrakk> x \<in> carrier G; y \<in> carrier G \<rbrakk> \<Longrightarrow> x \<oplus> ((\<ominus> x) \<oplus> y) = y" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 357 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 358 | assume G: "x \<in> carrier G" "y \<in> carrier G" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 359 | then have "(x \<oplus> \<ominus> x) \<oplus> y = y" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 360 | by (simp only: r_neg l_zero) | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 361 | with G show ?thesis | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 362 | by (simp add: a_ac) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 363 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 364 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 365 | context ring begin | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 366 | |
| 61382 | 367 | text \<open> | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 368 | The following proofs are from Jacobson, Basic Algebra I, pp.~88--89. | 
| 61382 | 369 | \<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 370 | |
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 371 | sublocale semiring | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 372 | proof - | 
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 373 | note [simp] = ring_axioms[unfolded ring_def ring_axioms_def] | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 374 | show "semiring R" | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 375 | proof (unfold_locales) | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 376 | fix x | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 377 | assume R: "x \<in> carrier R" | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 378 | 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: 
58811diff
changeset | 379 | by (simp del: l_zero r_zero) | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 380 | also from R have "... = \<zero> \<otimes> x \<oplus> \<zero>" by simp | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 381 | 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: 
58811diff
changeset | 382 | 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: 
58811diff
changeset | 383 | 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: 
58811diff
changeset | 384 | by (simp del: l_zero r_zero) | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 385 | also from R have "... = x \<otimes> \<zero> \<oplus> \<zero>" by simp | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 386 | 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: 
58811diff
changeset | 387 | 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: 
58811diff
changeset | 388 | qed auto | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 389 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 390 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 391 | lemma l_minus: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 392 | "\<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> (\<ominus> x) \<otimes> y = \<ominus> (x \<otimes> y)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 393 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 394 | assume R: "x \<in> carrier R" "y \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 395 | then have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = (\<ominus> x \<oplus> x) \<otimes> y" by (simp add: l_distr) | 
| 44677 | 396 | also from R have "... = \<zero>" by (simp add: l_neg) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 397 | finally have "(\<ominus> x) \<otimes> y \<oplus> x \<otimes> y = \<zero>" . | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 398 | 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: 
20318diff
changeset | 399 | 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 | 400 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 401 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 402 | lemma r_minus: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 403 | "\<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> x \<otimes> (\<ominus> y) = \<ominus> (x \<otimes> y)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 404 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 405 | assume R: "x \<in> carrier R" "y \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 406 | then have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = x \<otimes> (\<ominus> y \<oplus> y)" by (simp add: r_distr) | 
| 44677 | 407 | also from R have "... = \<zero>" by (simp add: l_neg) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 408 | finally have "x \<otimes> (\<ominus> y) \<oplus> x \<otimes> y = \<zero>" . | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 409 | 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 | 410 | with R show ?thesis by (simp add: a_assoc r_neg ) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 411 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 412 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 413 | end | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 414 | |
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 415 | lemma (in abelian_group) minus_eq: "x \<ominus> y = x \<oplus> (\<ominus> y)" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 416 | by (rule a_minus_def) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 417 | |
| 61382 | 418 | text \<open>Setup algebra method: | 
| 419 | compute distributive normal form in locale contexts\<close> | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 420 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 421 | |
| 69605 | 422 | ML_file \<open>ringsimp.ML\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 423 | |
| 61382 | 424 | attribute_setup algebra = \<open> | 
| 58811 | 425 | Scan.lift ((Args.add >> K true || Args.del >> K false) --| Args.colon || Scan.succeed true) | 
| 426 | -- Scan.lift Args.name -- Scan.repeat Args.term | |
| 427 | >> (fn ((b, n), ts) => if b then Ringsimp.add_struct (n, ts) else Ringsimp.del_struct (n, ts)) | |
| 61382 | 428 | \<close> "theorems controlling algebra method" | 
| 47701 | 429 | |
| 61382 | 430 | method_setup algebra = \<open> | 
| 58811 | 431 | Scan.succeed (SIMPLE_METHOD' o Ringsimp.algebra_tac) | 
| 61382 | 432 | \<close> "normalisation of algebraic structure" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 433 | |
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 434 | lemmas (in semiring) semiring_simprules | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 435 | [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: 
58811diff
changeset | 436 | a_closed zero_closed m_closed one_closed | 
| 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 437 | a_assoc l_zero a_comm m_assoc l_one l_distr r_zero | 
| 68517 | 438 | a_lcomm r_distr l_null r_null | 
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 439 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 440 | lemmas (in ring) ring_simprules | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 441 | [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 | 442 | 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 | 443 | 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 | 444 | 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 | 445 | 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 | 446 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 447 | lemmas (in cring) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 448 | [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 | 449 | _ | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 450 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 451 | lemmas (in cring) cring_simprules | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 452 | [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 | 453 | 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 | 454 | 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 | 455 | 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 | 456 | 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 | 457 | |
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 458 | lemma (in semiring) nat_pow_zero: | 
| 67341 
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
 nipkow parents: 
67091diff
changeset | 459 | "(n::nat) \<noteq> 0 \<Longrightarrow> \<zero> [^] n = \<zero>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 460 | by (induct n) simp_all | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 461 | |
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 462 | context semiring begin | 
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 463 | |
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 464 | lemma one_zeroD: | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 465 | assumes onezero: "\<one> = \<zero>" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 466 |   shows "carrier R = {\<zero>}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 467 | proof (rule, rule) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 468 | fix x | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 469 | assume xcarr: "x \<in> carrier R" | 
| 47409 | 470 | from xcarr have "x = x \<otimes> \<one>" by simp | 
| 471 | with onezero have "x = x \<otimes> \<zero>" by simp | |
| 472 | with xcarr have "x = \<zero>" by simp | |
| 473 |   then show "x \<in> {\<zero>}" by fast
 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 474 | qed fast | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 475 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 476 | lemma one_zeroI: | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 477 |   assumes carrzero: "carrier R = {\<zero>}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 478 | shows "\<one> = \<zero>" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 479 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 480 | from one_closed and carrzero | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 481 | show "\<one> = \<zero>" by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 482 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 483 | |
| 46721 | 484 | lemma carrier_one_zero: "(carrier R = {\<zero>}) = (\<one> = \<zero>)"
 | 
| 68673 | 485 | using one_zeroD by blast | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 486 | |
| 46721 | 487 | lemma carrier_one_not_zero: "(carrier R \<noteq> {\<zero>}) = (\<one> \<noteq> \<zero>)"
 | 
| 27717 
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
 ballarin parents: 
27714diff
changeset | 488 | by (simp add: carrier_one_zero) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 489 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 490 | end | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 491 | |
| 61382 | 492 | text \<open>Two examples for use of method algebra\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 493 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 494 | lemma | 
| 27611 | 495 | fixes R (structure) and S (structure) | 
| 496 | assumes "ring R" "cring S" | |
| 497 | assumes RS: "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier S" "d \<in> carrier S" | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 498 | shows "a \<oplus> (\<ominus> (a \<oplus> (\<ominus> b))) = b \<and> c \<otimes>\<^bsub>S\<^esub> d = d \<otimes>\<^bsub>S\<^esub> c" | 
| 27611 | 499 | proof - | 
| 29237 | 500 | interpret ring R by fact | 
| 501 | interpret cring S by fact | |
| 27611 | 502 | from RS show ?thesis by algebra | 
| 503 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 504 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 505 | lemma | 
| 27611 | 506 | fixes R (structure) | 
| 507 | assumes "ring R" | |
| 508 | assumes R: "a \<in> carrier R" "b \<in> carrier R" | |
| 509 | shows "a \<ominus> (a \<ominus> b) = b" | |
| 510 | proof - | |
| 29237 | 511 | interpret ring R by fact | 
| 27611 | 512 | from R show ?thesis by algebra | 
| 513 | qed | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 514 | |
| 35849 | 515 | |
| 61382 | 516 | subsubsection \<open>Sums over Finite Sets\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 517 | |
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 518 | lemma (in semiring) finsum_ldistr: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 519 | "\<lbrakk> finite A; a \<in> carrier R; f: A \<rightarrow> carrier R \<rbrakk> \<Longrightarrow> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 520 | (\<Oplus> i \<in> A. (f i)) \<otimes> a = (\<Oplus> i \<in> A. ((f i) \<otimes> a))" | 
| 22265 | 521 | proof (induct set: finite) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 522 | case empty then show ?case by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 523 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 524 | 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 | 525 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 526 | |
| 59851 
43b1870b3e61
added locale for semirings
 Rene Thiemann <rene.thiemann@uibk.ac.at> parents: 
58811diff
changeset | 527 | lemma (in semiring) finsum_rdistr: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 528 | "\<lbrakk> finite A; a \<in> carrier R; f: A \<rightarrow> carrier R \<rbrakk> \<Longrightarrow> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 529 | a \<otimes> (\<Oplus> i \<in> A. (f i)) = (\<Oplus> i \<in> A. (a \<otimes> (f i)))" | 
| 22265 | 530 | proof (induct set: finite) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 531 | case empty then show ?case by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 532 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 533 | 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 | 534 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 535 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 536 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 537 | text \<open>A quick detour\<close> | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 538 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
69605diff
changeset | 539 | lemma add_pow_int_ge: "(k :: int) \<ge> 0 \<Longrightarrow> [ k ] \<cdot>\<^bsub>R\<^esub> a = [ nat k ] \<cdot>\<^bsub>R\<^esub> a" \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 540 | by (simp add: add_pow_def int_pow_def nat_pow_def) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 541 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
69605diff
changeset | 542 | lemma add_pow_int_lt: "(k :: int) < 0 \<Longrightarrow> [ k ] \<cdot>\<^bsub>R\<^esub> a = \<ominus>\<^bsub>R\<^esub> ([ nat (- k) ] \<cdot>\<^bsub>R\<^esub> a)" \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68517 | 543 | by (simp add: int_pow_def nat_pow_def a_inv_def add_pow_def) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 544 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
69605diff
changeset | 545 | corollary (in semiring) add_pow_ldistr: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 546 | assumes "a \<in> carrier R" "b \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 547 | shows "([(k :: nat)] \<cdot> a) \<otimes> b = [k] \<cdot> (a \<otimes> b)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 548 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 549 |   have "([k] \<cdot> a) \<otimes> b = (\<Oplus> i \<in> {..< k}. a) \<otimes> b"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 550 |     using add.finprod_const[OF assms(1), of "{..<k}"] by simp
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 551 |   also have " ... = (\<Oplus> i \<in> {..< k}. (a \<otimes> b))"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 552 |     using finsum_ldistr[of "{..<k}" b "\<lambda>x. a"] assms by simp
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 553 | also have " ... = [k] \<cdot> (a \<otimes> b)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 554 |     using add.finprod_const[of "a \<otimes> b" "{..<k}"] assms by simp
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 555 | finally show ?thesis . | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 556 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 557 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
69605diff
changeset | 558 | corollary (in semiring) add_pow_rdistr: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 559 | assumes "a \<in> carrier R" "b \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 560 | shows "a \<otimes> ([(k :: nat)] \<cdot> b) = [k] \<cdot> (a \<otimes> b)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 561 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 562 |   have "a \<otimes> ([k] \<cdot> b) = a \<otimes> (\<Oplus> i \<in> {..< k}. b)"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 563 |     using add.finprod_const[OF assms(2), of "{..<k}"] by simp
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 564 |   also have " ... = (\<Oplus> i \<in> {..< k}. (a \<otimes> b))"
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 565 |     using finsum_rdistr[of "{..<k}" a "\<lambda>x. b"] assms by simp
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 566 | also have " ... = [k] \<cdot> (a \<otimes> b)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 567 |     using add.finprod_const[of "a \<otimes> b" "{..<k}"] assms by simp
 | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 568 | finally show ?thesis . | 
| 68517 | 569 | qed | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 570 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 571 | (* For integers, we need the uniqueness of the additive inverse *) | 
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
69605diff
changeset | 572 | lemma (in ring) add_pow_ldistr_int: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 573 | assumes "a \<in> carrier R" "b \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 574 | shows "([(k :: int)] \<cdot> a) \<otimes> b = [k] \<cdot> (a \<otimes> b)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 575 | proof (cases "k \<ge> 0") | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 576 | case True thus ?thesis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 577 | using add_pow_int_ge[of k R] add_pow_ldistr[OF assms] by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 578 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 579 | case False thus ?thesis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 580 | using add_pow_int_lt[of k R a] add_pow_int_lt[of k R "a \<otimes> b"] | 
| 68517 | 581 | add_pow_ldistr[OF assms, of "nat (- k)"] assms l_minus by auto | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 582 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 583 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
69605diff
changeset | 584 | lemma (in ring) add_pow_rdistr_int: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 585 | assumes "a \<in> carrier R" "b \<in> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 586 | shows "a \<otimes> ([(k :: int)] \<cdot> b) = [k] \<cdot> (a \<otimes> b)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 587 | proof (cases "k \<ge> 0") | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 588 | case True thus ?thesis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 589 | using add_pow_int_ge[of k R] add_pow_rdistr[OF assms] by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 590 | next | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 591 | case False thus ?thesis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 592 | using add_pow_int_lt[of k R b] add_pow_int_lt[of k R "a \<otimes> b"] | 
| 68517 | 593 | add_pow_rdistr[OF assms, of "nat (- k)"] assms r_minus by auto | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 594 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 595 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 596 | |
| 61382 | 597 | subsection \<open>Integral Domains\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 598 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 599 | context "domain" begin | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 600 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 601 | lemma zero_not_one [simp]: "\<zero> \<noteq> \<one>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 602 | by (rule not_sym) simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 603 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 604 | lemma integral_iff: (* not by default a simp rule! *) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 605 | "\<lbrakk> a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> (a \<otimes> b = \<zero>) = (a = \<zero> \<or> b = \<zero>)" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 606 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 607 | assume "a \<in> carrier R" "b \<in> carrier R" "a \<otimes> b = \<zero>" | 
| 67091 | 608 | then show "a = \<zero> \<or> b = \<zero>" by (simp add: integral) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 609 | next | 
| 67091 | 610 | assume "a \<in> carrier R" "b \<in> carrier R" "a = \<zero> \<or> b = \<zero>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 611 | then show "a \<otimes> b = \<zero>" by auto | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 612 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 613 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 614 | lemma m_lcancel: | 
| 67091 | 615 | assumes prem: "a \<noteq> \<zero>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 616 | 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 | 617 | shows "(a \<otimes> b = a \<otimes> c) = (b = c)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 618 | proof | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 619 | assume eq: "a \<otimes> b = a \<otimes> c" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 620 | with R have "a \<otimes> (b \<ominus> c) = \<zero>" by algebra | 
| 67091 | 621 | with R have "a = \<zero> \<or> (b \<ominus> c) = \<zero>" by (simp add: integral_iff) | 
| 68517 | 622 | with prem and R have "b \<ominus> c = \<zero>" by auto | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 623 | with R have "b = b \<ominus> (b \<ominus> c)" by algebra | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 624 | 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 | 625 | finally show "b = c" . | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 626 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 627 | 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 | 628 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 629 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 630 | lemma m_rcancel: | 
| 67091 | 631 | assumes prem: "a \<noteq> \<zero>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 632 | 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 | 633 | shows conc: "(b \<otimes> a = c \<otimes> a) = (b = c)" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 634 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 635 | 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 | 636 | with R show ?thesis by algebra | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 637 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 638 | |
| 41433 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 639 | end | 
| 
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
 ballarin parents: 
35849diff
changeset | 640 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 641 | |
| 61382 | 642 | subsection \<open>Fields\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 643 | |
| 61382 | 644 | text \<open>Field would not need to be derived from domain, the properties | 
| 645 | for domain follow from the assumptions of field\<close> | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 646 | |
| 70214 
58191e01f0b1
moving around some material from Algebraic_Closure
 paulson <lp15@cam.ac.uk> parents: 
69895diff
changeset | 647 | lemma (in field) is_ring: "ring R" | 
| 
58191e01f0b1
moving around some material from Algebraic_Closure
 paulson <lp15@cam.ac.uk> parents: 
69895diff
changeset | 648 | using ring_axioms . | 
| 
58191e01f0b1
moving around some material from Algebraic_Closure
 paulson <lp15@cam.ac.uk> parents: 
69895diff
changeset | 649 | |
| 68551 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 650 | lemma fieldE : | 
| 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 651 | fixes R (structure) | 
| 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 652 | assumes "field R" | 
| 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 653 | shows "cring R" | 
| 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 654 | and one_not_zero : "\<one> \<noteq> \<zero>" | 
| 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 655 | and integral: "\<And>a b. \<lbrakk> a \<otimes> b = \<zero>; a \<in> carrier R; b \<in> carrier R \<rbrakk> \<Longrightarrow> a = \<zero> \<or> b = \<zero>" | 
| 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 656 |   and field_Units: "Units R = carrier R - {\<zero>}"
 | 
| 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 657 | using assms unfolding field_def field_axioms_def domain_def domain_axioms_def by simp_all | 
| 
b680e74eb6f2
More on Algebra by Paulo and Martin
 paulson <lp15@cam.ac.uk> parents: 
68517diff
changeset | 658 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 659 | lemma (in cring) cring_fieldI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 660 |   assumes field_Units: "Units R = carrier R - {\<zero>}"
 | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 661 | shows "field R" | 
| 28823 | 662 | proof | 
| 47409 | 663 | from field_Units have "\<zero> \<notin> Units R" by fast | 
| 664 | moreover have "\<one> \<in> Units R" by fast | |
| 665 | ultimately show "\<one> \<noteq> \<zero>" by force | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 666 | next | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 667 | fix a b | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 668 | assume acarr: "a \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 669 | and bcarr: "b \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 670 | and ab: "a \<otimes> b = \<zero>" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 671 | show "a = \<zero> \<or> b = \<zero>" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 672 | proof (cases "a = \<zero>", simp) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 673 | assume "a \<noteq> \<zero>" | 
| 47409 | 674 | with field_Units and acarr have aUnit: "a \<in> Units R" by fast | 
| 675 | from bcarr have "b = \<one> \<otimes> b" by algebra | |
| 676 | 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 | 677 | also from acarr bcarr aUnit[THEN Units_inv_closed] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 678 | have "... = (inv a) \<otimes> (a \<otimes> b)" by algebra | 
| 47409 | 679 | also from ab and acarr bcarr aUnit have "... = (inv a) \<otimes> \<zero>" by simp | 
| 680 | also from aUnit[THEN Units_inv_closed] have "... = \<zero>" by algebra | |
| 681 | finally have "b = \<zero>" . | |
| 682 | then show "a = \<zero> \<or> b = \<zero>" by simp | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 683 | qed | 
| 23350 | 684 | qed (rule field_Units) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 685 | |
| 61382 | 686 | text \<open>Another variant to show that something is a field\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 687 | lemma (in cring) cring_fieldI2: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 688 | assumes notzero: "\<zero> \<noteq> \<one>" | 
| 68673 | 689 | and invex: "\<And>a. \<lbrakk>a \<in> carrier R; a \<noteq> \<zero>\<rbrakk> \<Longrightarrow> \<exists>b\<in>carrier R. a \<otimes> b = \<one>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 690 | shows "field R" | 
| 68673 | 691 | proof - | 
| 692 |   have *: "carrier R - {\<zero>} \<subseteq> {y \<in> carrier R. \<exists>x\<in>carrier R. x \<otimes> y = \<one> \<and> y \<otimes> x = \<one>}"
 | |
| 693 | proof (clarsimp) | |
| 694 | fix x | |
| 695 | assume xcarr: "x \<in> carrier R" and "x \<noteq> \<zero>" | |
| 696 | obtain y where ycarr: "y \<in> carrier R" and xy: "x \<otimes> y = \<one>" | |
| 697 | using \<open>x \<noteq> \<zero>\<close> invex xcarr by blast | |
| 698 | with ycarr and xy show "\<exists>y\<in>carrier R. y \<otimes> x = \<one> \<and> x \<otimes> y = \<one>" | |
| 699 | using m_comm xcarr by fastforce | |
| 700 | qed | |
| 701 | show ?thesis | |
| 702 | apply (rule cring_fieldI, simp add: Units_def) | |
| 703 | using * | |
| 704 | using group_l_invI notzero set_diff_eq by auto | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 705 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 706 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 707 | |
| 61382 | 708 | subsection \<open>Morphisms\<close> | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 709 | |
| 35847 | 710 | definition | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 711 |   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: 
35847diff
changeset | 712 | where "ring_hom R S = | 
| 67091 | 713 |     {h. h \<in> carrier R \<rightarrow> carrier S \<and>
 | 
| 714 | (\<forall>x y. x \<in> carrier R \<and> y \<in> carrier R \<longrightarrow> | |
| 715 | h (x \<otimes>\<^bsub>R\<^esub> y) = h x \<otimes>\<^bsub>S\<^esub> h y \<and> h (x \<oplus>\<^bsub>R\<^esub> y) = h x \<oplus>\<^bsub>S\<^esub> h y) \<and> | |
| 35847 | 716 | h \<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>}" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 717 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 718 | lemma ring_hom_memI: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 719 | fixes R (structure) and S (structure) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 720 | assumes "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 721 | and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 722 | and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 723 | and "h \<one> = \<one>\<^bsub>S\<^esub>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 724 | shows "h \<in> ring_hom R S" | 
| 27714 
27b4d7c01f8b
Tuned (for the sake of a meaningless log entry).
 ballarin parents: 
27699diff
changeset | 725 | by (auto simp add: ring_hom_def assms Pi_def) | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 726 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 727 | lemma ring_hom_memE: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 728 | fixes R (structure) and S (structure) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 729 | assumes "h \<in> ring_hom R S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 730 | shows "\<And>x. x \<in> carrier R \<Longrightarrow> h x \<in> carrier S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 731 | and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 732 | and "\<And>x y. \<lbrakk> x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 733 | and "h \<one> = \<one>\<^bsub>S\<^esub>" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 734 | using assms unfolding ring_hom_def by auto | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 735 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 736 | lemma ring_hom_closed: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 737 | "\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R \<rbrakk> \<Longrightarrow> h x \<in> carrier S" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 738 | by (auto simp add: ring_hom_def funcset_mem) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 739 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 740 | lemma ring_hom_mult: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 741 | fixes R (structure) and S (structure) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 742 | shows "\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<otimes> y) = h x \<otimes>\<^bsub>S\<^esub> h y" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 743 | by (simp add: ring_hom_def) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 744 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 745 | lemma ring_hom_add: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 746 | fixes R (structure) and S (structure) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 747 | shows "\<lbrakk> h \<in> ring_hom R S; x \<in> carrier R; y \<in> carrier R \<rbrakk> \<Longrightarrow> h (x \<oplus> y) = h x \<oplus>\<^bsub>S\<^esub> h y" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 748 | by (simp add: ring_hom_def) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 749 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 750 | lemma ring_hom_one: | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 751 | fixes R (structure) and S (structure) | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 752 | shows "h \<in> ring_hom R S \<Longrightarrow> h \<one> = \<one>\<^bsub>S\<^esub>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 753 | by (simp add: ring_hom_def) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 754 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 755 | lemma ring_hom_zero: | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 756 | fixes R (structure) and S (structure) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 757 | assumes "h \<in> ring_hom R S" "ring R" "ring S" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 758 | shows "h \<zero> = \<zero>\<^bsub>S\<^esub>" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 759 | proof - | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 760 | have "h \<zero> = h \<zero> \<oplus>\<^bsub>S\<^esub> h \<zero>" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 761 | using ring_hom_add[OF assms(1), of \<zero> \<zero>] assms(2) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 762 | by (simp add: ring.ring_simprules(2) ring.ring_simprules(15)) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 763 | thus ?thesis | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 764 | by (metis abelian_group.l_neg assms ring.is_abelian_group ring.ring_simprules(18) ring.ring_simprules(2) ring_hom_closed) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 765 | qed | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 766 | |
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 767 | locale ring_hom_cring = | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 768 | R?: cring R + S?: cring S for R (structure) and S (structure) + fixes h | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 769 | assumes homh [simp, intro]: "h \<in> ring_hom R S" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 770 | notes hom_closed [simp, intro] = ring_hom_closed [OF homh] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 771 | and hom_mult [simp] = ring_hom_mult [OF homh] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 772 | and hom_add [simp] = ring_hom_add [OF homh] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 773 | and hom_one [simp] = ring_hom_one [OF homh] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 774 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 775 | lemma (in ring_hom_cring) hom_zero [simp]: "h \<zero> = \<zero>\<^bsub>S\<^esub>" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 776 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 777 | 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 | 778 | by (simp add: hom_add [symmetric] del: hom_add) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 779 | then show ?thesis by (simp del: S.r_zero) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 780 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 781 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 782 | lemma (in ring_hom_cring) hom_a_inv [simp]: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 783 | "x \<in> carrier R \<Longrightarrow> h (\<ominus> x) = \<ominus>\<^bsub>S\<^esub> h x" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 784 | proof - | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 785 | assume R: "x \<in> carrier R" | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 786 | 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 | 787 | 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 | 788 | with R show ?thesis by simp | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 789 | qed | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 790 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 791 | lemma (in ring_hom_cring) hom_finsum [simp]: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 792 | assumes "f: A \<rightarrow> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 793 | shows "h (\<Oplus> i \<in> A. f i) = (\<Oplus>\<^bsub>S\<^esub> i \<in> A. (h o f) i)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 794 | using assms 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 | 795 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 796 | lemma (in ring_hom_cring) hom_finprod: | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 797 | assumes "f: A \<rightarrow> carrier R" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 798 | shows "h (\<Otimes> i \<in> A. f i) = (\<Otimes>\<^bsub>S\<^esub> i \<in> A. (h o f) i)" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 799 | using assms 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 | 800 | |
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 801 | declare ring_hom_cring.hom_finprod [simp] | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 802 | |
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 803 | lemma id_ring_hom [simp]: "id \<in> ring_hom R R" | 
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 804 | by (auto intro!: ring_hom_memI) | 
| 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 805 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
69605diff
changeset | 806 | lemma ring_hom_trans: \<^marker>\<open>contributor \<open>Paulo Emílio de Vilhena\<close>\<close> | 
| 68443 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 807 | "\<lbrakk> f \<in> ring_hom R S; g \<in> ring_hom S T \<rbrakk> \<Longrightarrow> g \<circ> f \<in> ring_hom R T" | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 808 | by (rule ring_hom_memI) (auto simp add: ring_hom_closed ring_hom_mult ring_hom_add ring_hom_one) | 
| 
43055b016688
New material from Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68399diff
changeset | 809 | |
| 69272 | 810 | subsection\<open>Jeremy Avigad's \<open>More_Finite_Product\<close> material\<close> | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 811 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 812 | (* need better simplification rules for rings *) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 813 | (* the next one holds more generally for abelian groups *) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 814 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 815 | lemma (in cring) sum_zero_eq_neg: "x \<in> carrier R \<Longrightarrow> y \<in> carrier R \<Longrightarrow> x \<oplus> y = \<zero> \<Longrightarrow> x = \<ominus> y" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 816 | by (metis minus_equality) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 817 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 818 | lemma (in domain) square_eq_one: | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 819 | fixes x | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 820 | assumes [simp]: "x \<in> carrier R" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 821 | and "x \<otimes> x = \<one>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 822 | shows "x = \<one> \<or> x = \<ominus>\<one>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 823 | proof - | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 824 | have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = x \<otimes> x \<oplus> \<ominus> \<one>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 825 | by (simp add: ring_simprules) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 826 | also from \<open>x \<otimes> x = \<one>\<close> have "\<dots> = \<zero>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 827 | by (simp add: ring_simprules) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 828 | finally have "(x \<oplus> \<one>) \<otimes> (x \<oplus> \<ominus> \<one>) = \<zero>" . | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 829 | then have "(x \<oplus> \<one>) = \<zero> \<or> (x \<oplus> \<ominus> \<one>) = \<zero>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 830 | by (intro integral) auto | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 831 | then show ?thesis | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 832 | by (metis add.inv_closed add.inv_solve_right assms(1) l_zero one_closed zero_closed) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 833 | qed | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 834 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 835 | lemma (in domain) inv_eq_self: "x \<in> Units R \<Longrightarrow> x = inv x \<Longrightarrow> x = \<one> \<or> x = \<ominus>\<one>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 836 | by (metis Units_closed Units_l_inv square_eq_one) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 837 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 838 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 839 | text \<open> | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 840 | The following translates theorems about groups to the facts about | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 841 | the units of a ring. (The list should be expanded as more things are | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 842 | needed.) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 843 | \<close> | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 844 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 845 | lemma (in ring) finite_ring_finite_units [intro]: "finite (carrier R) \<Longrightarrow> finite (Units R)" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 846 | by (rule finite_subset) auto | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 847 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 848 | lemma (in monoid) units_of_pow: | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 849 | fixes n :: nat | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 850 | shows "x \<in> Units G \<Longrightarrow> x [^]\<^bsub>units_of G\<^esub> n = x [^]\<^bsub>G\<^esub> n" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 851 | apply (induct n) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 852 | apply (auto simp add: units_group group.is_monoid | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 853 | monoid.nat_pow_0 monoid.nat_pow_Suc units_of_one units_of_mult) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 854 | done | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 855 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 856 | lemma (in cring) units_power_order_eq_one: | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 857 | "finite (Units R) \<Longrightarrow> a \<in> Units R \<Longrightarrow> a [^] card(Units R) = \<one>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 858 | by (metis comm_group.power_order_eq_one units_comm_group units_of_carrier units_of_one units_of_pow) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 859 | |
| 69272 | 860 | subsection\<open>Jeremy Avigad's \<open>More_Ring\<close> material\<close> | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 861 | |
| 68673 | 862 | lemma (in cring) field_intro2: | 
| 863 |   assumes "\<zero>\<^bsub>R\<^esub> \<noteq> \<one>\<^bsub>R\<^esub>" and un: "\<And>x. x \<in> carrier R - {\<zero>\<^bsub>R\<^esub>} \<Longrightarrow> x \<in> Units R"
 | |
| 864 | shows "field R" | |
| 865 | proof unfold_locales | |
| 866 | show "\<one> \<noteq> \<zero>" using assms by auto | |
| 867 | show "\<lbrakk>a \<otimes> b = \<zero>; a \<in> carrier R; | |
| 868 | b \<in> carrier R\<rbrakk> | |
| 869 | \<Longrightarrow> a = \<zero> \<or> b = \<zero>" for a b | |
| 870 | by (metis un Units_l_cancel insert_Diff_single insert_iff r_null zero_closed) | |
| 871 | qed (use assms in \<open>auto simp: Units_def\<close>) | |
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 872 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 873 | lemma (in monoid) inv_char: | 
| 68673 | 874 | assumes "x \<in> carrier G" "y \<in> carrier G" "x \<otimes> y = \<one>" "y \<otimes> x = \<one>" | 
| 875 | shows "inv x = y" | |
| 876 | using assms inv_unique' by auto | |
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 877 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 878 | lemma (in comm_monoid) comm_inv_char: "x \<in> carrier G \<Longrightarrow> y \<in> carrier G \<Longrightarrow> x \<otimes> y = \<one> \<Longrightarrow> inv x = y" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 879 | by (simp add: inv_char m_comm) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 880 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 881 | lemma (in ring) inv_neg_one [simp]: "inv (\<ominus> \<one>) = \<ominus> \<one>" | 
| 68673 | 882 | by (simp add: inv_char local.ring_axioms ring.r_minus) | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 883 | |
| 78522 
918a9ed06898
some refinements in Algebra and Number_Theory
 paulson <lp15@cam.ac.uk> parents: 
75963diff
changeset | 884 | lemma (in monoid) inv_eq_imp_eq [dest!]: "inv x = inv y \<Longrightarrow> x \<in> Units G \<Longrightarrow> y \<in> Units G \<Longrightarrow> x = y" | 
| 68673 | 885 | by (metis Units_inv_inv) | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 886 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 887 | lemma (in ring) Units_minus_one_closed [intro]: "\<ominus> \<one> \<in> Units R" | 
| 68673 | 888 | by (simp add: Units_def) (metis add.l_inv_ex local.minus_minus minus_equality one_closed r_minus r_one) | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 889 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 890 | lemma (in ring) inv_eq_neg_one_eq: "x \<in> Units R \<Longrightarrow> inv x = \<ominus> \<one> \<longleftrightarrow> x = \<ominus> \<one>" | 
| 68673 | 891 | by (metis Units_inv_inv inv_neg_one) | 
| 68445 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 892 | |
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 893 | lemma (in monoid) inv_eq_one_eq: "x \<in> Units G \<Longrightarrow> inv x = \<one> \<longleftrightarrow> x = \<one>" | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 894 | by (metis Units_inv_inv inv_one) | 
| 
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
 paulson <lp15@cam.ac.uk> parents: 
68443diff
changeset | 895 | |
| 20318 
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
 ballarin parents: diff
changeset | 896 | end |