Up to index of Isabelle/HOL/HOL-Algebra
theory Ring(* Title: The algebraic hierarchy of rings Author: Clemens Ballarin, started 9 December 1996 Copyright: Clemens Ballarin *) theory Ring imports FiniteProduct uses ("ringsimp.ML") begin section {* The Algebraic Hierarchy of Rings *} subsection {* Abelian Groups *} record 'a ring = "'a monoid" + zero :: 'a ("\<zero>\<index>") add :: "['a, 'a] => 'a" (infixl "⊕\<index>" 65) text {* Derived operations. *} constdefs (structure R) a_inv :: "[('a, 'm) ring_scheme, 'a ] => 'a" ("\<ominus>\<index> _" [81] 80) "a_inv R == m_inv (| carrier = carrier R, mult = add R, one = zero R |)" a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" (infixl "\<ominus>\<index>" 65) "[| x ∈ carrier R; y ∈ carrier R |] ==> x \<ominus> y == x ⊕ (\<ominus> y)" locale abelian_monoid = fixes G (structure) assumes a_comm_monoid: "comm_monoid (| carrier = carrier G, mult = add G, one = zero G |)" text {* The following definition is redundant but simple to use. *} locale abelian_group = abelian_monoid + assumes a_comm_group: "comm_group (| carrier = carrier G, mult = add G, one = zero G |)" subsection {* Basic Properties *} lemma abelian_monoidI: fixes R (structure) assumes a_closed: "!!x y. [| x ∈ carrier R; y ∈ carrier R |] ==> x ⊕ y ∈ carrier R" and zero_closed: "\<zero> ∈ carrier R" and a_assoc: "!!x y z. [| x ∈ carrier R; y ∈ carrier R; z ∈ carrier R |] ==> (x ⊕ y) ⊕ z = x ⊕ (y ⊕ z)" and l_zero: "!!x. x ∈ carrier R ==> \<zero> ⊕ x = x" and a_comm: "!!x y. [| x ∈ carrier R; y ∈ carrier R |] ==> x ⊕ y = y ⊕ x" shows "abelian_monoid R" by (auto intro!: abelian_monoid.intro comm_monoidI intro: assms) lemma abelian_groupI: fixes R (structure) assumes a_closed: "!!x y. [| x ∈ carrier R; y ∈ carrier R |] ==> x ⊕ y ∈ carrier R" and zero_closed: "zero R ∈ carrier R" and a_assoc: "!!x y z. [| x ∈ carrier R; y ∈ carrier R; z ∈ carrier R |] ==> (x ⊕ y) ⊕ z = x ⊕ (y ⊕ z)" and a_comm: "!!x y. [| x ∈ carrier R; y ∈ carrier R |] ==> x ⊕ y = y ⊕ x" and l_zero: "!!x. x ∈ carrier R ==> \<zero> ⊕ x = x" and l_inv_ex: "!!x. x ∈ carrier R ==> EX y : carrier R. y ⊕ x = \<zero>" shows "abelian_group R" by (auto intro!: abelian_group.intro abelian_monoidI abelian_group_axioms.intro comm_monoidI comm_groupI intro: assms) lemma (in abelian_monoid) a_monoid: "monoid (| carrier = carrier G, mult = add G, one = zero G |)" by (rule comm_monoid.axioms, rule a_comm_monoid) lemma (in abelian_group) a_group: "group (| carrier = carrier G, mult = add G, one = zero G |)" by (simp add: group_def a_monoid) (simp add: comm_group.axioms group.axioms a_comm_group) lemmas monoid_record_simps = partial_object.simps monoid.simps lemma (in abelian_monoid) a_closed [intro, simp]: "[| x ∈ carrier G; y ∈ carrier G |] ==> x ⊕ y ∈ carrier G" by (rule monoid.m_closed [OF a_monoid, simplified monoid_record_simps]) lemma (in abelian_monoid) zero_closed [intro, simp]: "\<zero> ∈ carrier G" by (rule monoid.one_closed [OF a_monoid, simplified monoid_record_simps]) lemma (in abelian_group) a_inv_closed [intro, simp]: "x ∈ carrier G ==> \<ominus> x ∈ carrier G" by (simp add: a_inv_def group.inv_closed [OF a_group, simplified monoid_record_simps]) lemma (in abelian_group) minus_closed [intro, simp]: "[| x ∈ carrier G; y ∈ carrier G |] ==> x \<ominus> y ∈ carrier G" by (simp add: a_minus_def) lemma (in abelian_group) a_l_cancel [simp]: "[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (x ⊕ y = x ⊕ z) = (y = z)" by (rule group.l_cancel [OF a_group, simplified monoid_record_simps]) lemma (in abelian_group) a_r_cancel [simp]: "[| x ∈ carrier G; y ∈ carrier G; z ∈ carrier G |] ==> (y ⊕ x = z ⊕ x) = (y = z)" by (rule group.r_cancel [OF a_group, simplified monoid_record_simps]) lemma (in abelian_monoid) a_assoc: "[|x ∈ carrier G; y ∈ carrier G; z ∈ carrier G|] ==> (x ⊕ y) ⊕ z = x ⊕ (y ⊕ z)" by (rule monoid.m_assoc [OF a_monoid, simplified monoid_record_simps]) lemma (in abelian_monoid) l_zero [simp]: "x ∈ carrier G ==> \<zero> ⊕ x = x" by (rule monoid.l_one [OF a_monoid, simplified monoid_record_simps]) lemma (in abelian_group) l_neg: "x ∈ carrier G ==> \<ominus> x ⊕ x = \<zero>" by (simp add: a_inv_def group.l_inv [OF a_group, simplified monoid_record_simps]) lemma (in abelian_monoid) a_comm: "[|x ∈ carrier G; y ∈ carrier G|] ==> x ⊕ y = y ⊕ x" by (rule comm_monoid.m_comm [OF a_comm_monoid, simplified monoid_record_simps]) lemma (in abelian_monoid) a_lcomm: "[|x ∈ carrier G; y ∈ carrier G; z ∈ carrier G|] ==> x ⊕ (y ⊕ z) = y ⊕ (x ⊕ z)" by (rule comm_monoid.m_lcomm [OF a_comm_monoid, simplified monoid_record_simps]) lemma (in abelian_monoid) r_zero [simp]: "x ∈ carrier G ==> x ⊕ \<zero> = x" using monoid.r_one [OF a_monoid] by simp lemma (in abelian_group) r_neg: "x ∈ carrier G ==> x ⊕ (\<ominus> x) = \<zero>" using group.r_inv [OF a_group] by (simp add: a_inv_def) lemma (in abelian_group) minus_zero [simp]: "\<ominus> \<zero> = \<zero>" by (simp add: a_inv_def group.inv_one [OF a_group, simplified monoid_record_simps]) lemma (in abelian_group) minus_minus [simp]: "x ∈ carrier G ==> \<ominus> (\<ominus> x) = x" using group.inv_inv [OF a_group, simplified monoid_record_simps] by (simp add: a_inv_def) lemma (in abelian_group) a_inv_inj: "inj_on (a_inv G) (carrier G)" using group.inv_inj [OF a_group, simplified monoid_record_simps] by (simp add: a_inv_def) lemma (in abelian_group) minus_add: "[| x ∈ carrier G; y ∈ carrier G |] ==> \<ominus> (x ⊕ y) = \<ominus> x ⊕ \<ominus> y" using comm_group.inv_mult [OF a_comm_group] by (simp add: a_inv_def) lemma (in abelian_group) minus_equality: "[| x ∈ carrier G; y ∈ carrier G; y ⊕ x = \<zero> |] ==> \<ominus> x = y" using group.inv_equality [OF a_group] by (auto simp add: a_inv_def) lemma (in abelian_monoid) minus_unique: "[| x ∈ carrier G; y ∈ carrier G; y' ∈ carrier G; y ⊕ x = \<zero>; x ⊕ y' = \<zero> |] ==> y = y'" using monoid.inv_unique [OF a_monoid] by (simp add: a_inv_def) lemmas (in abelian_monoid) a_ac = a_assoc a_comm a_lcomm text {* Derive an @{text "abelian_group"} from a @{text "comm_group"} *} lemma comm_group_abelian_groupI: fixes G (structure) assumes cg: "comm_group (|carrier = carrier G, mult = add G, one = zero G|))," shows "abelian_group G" proof - interpret comm_group "(|carrier = carrier G, mult = add G, one = zero G|))," by (rule cg) show "abelian_group G" .. qed subsection {* Sums over Finite Sets *} text {* This definition makes it easy to lift lemmas from @{term finprod}. *} constdefs finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b" "finsum G f A == finprod (| carrier = carrier G, mult = add G, one = zero G |) f A" syntax "_finsum" :: "index => idt => 'a set => 'b => 'b" ("(3\<Oplus>__:_. _)" [1000, 0, 51, 10] 10) syntax (xsymbols) "_finsum" :: "index => idt => 'a set => 'b => 'b" ("(3\<Oplus>__∈_. _)" [1000, 0, 51, 10] 10) syntax (HTML output) "_finsum" :: "index => idt => 'a set => 'b => 'b" ("(3\<Oplus>__∈_. _)" [1000, 0, 51, 10] 10) translations "\<Oplus>\<index>i:A. b" == "finsum \<struct>\<index> (%i. b) A" -- {* Beware of argument permutation! *} context abelian_monoid begin (* lemmas finsum_empty [simp] = comm_monoid.finprod_empty [OF a_comm_monoid, simplified] is dangeous, because attributes (like simplified) are applied upon opening the locale, simplified refers to the simpset at that time!!! lemmas finsum_empty [simp] = abelian_monoid.finprod_empty [OF a_abelian_monoid, folded finsum_def, simplified monoid_record_simps] makes the locale slow, because proofs are repeated for every "lemma (in abelian_monoid)" command. When lemma is used time in UnivPoly.thy from beginning to UP_cring goes down from 110 secs to 60 secs. *) lemma finsum_empty [simp]: "finsum G f {} = \<zero>" by (rule comm_monoid.finprod_empty [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_insert [simp]: "[| finite F; a ∉ F; f ∈ F -> carrier G; f a ∈ carrier G |] ==> finsum G f (insert a F) = f a ⊕ finsum G f F" by (rule comm_monoid.finprod_insert [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_zero [simp]: "finite A ==> (\<Oplus>i∈A. \<zero>) = \<zero>" by (rule comm_monoid.finprod_one [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_closed [simp]: fixes A assumes fin: "finite A" and f: "f ∈ A -> carrier G" shows "finsum G f A ∈ carrier G" apply (rule comm_monoid.finprod_closed [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) apply (rule fin) apply (rule f) done lemma finsum_Un_Int: "[| finite A; finite B; g ∈ A -> carrier G; g ∈ B -> carrier G |] ==> finsum G g (A Un B) ⊕ finsum G g (A Int B) = finsum G g A ⊕ finsum G g B" by (rule comm_monoid.finprod_Un_Int [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_Un_disjoint: "[| finite A; finite B; A Int B = {}; g ∈ A -> carrier G; g ∈ B -> carrier G |] ==> finsum G g (A Un B) = finsum G g A ⊕ finsum G g B" by (rule comm_monoid.finprod_Un_disjoint [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_addf: "[| finite A; f ∈ A -> carrier G; g ∈ A -> carrier G |] ==> finsum G (%x. f x ⊕ g x) A = (finsum G f A ⊕ finsum G g A)" by (rule comm_monoid.finprod_multf [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_cong': "[| A = B; g : B -> carrier G; !!i. i : B ==> f i = g i |] ==> finsum G f A = finsum G g B" by (rule comm_monoid.finprod_cong' [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) auto lemma finsum_0 [simp]: "f : {0::nat} -> carrier G ==> finsum G f {..0} = f 0" by (rule comm_monoid.finprod_0 [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_Suc [simp]: "f : {..Suc n} -> carrier G ==> finsum G f {..Suc n} = (f (Suc n) ⊕ finsum G f {..n})" by (rule comm_monoid.finprod_Suc [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_Suc2: "f : {..Suc n} -> carrier G ==> finsum G f {..Suc n} = (finsum G (%i. f (Suc i)) {..n} ⊕ f 0)" by (rule comm_monoid.finprod_Suc2 [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_add [simp]: "[| f : {..n} -> carrier G; g : {..n} -> carrier G |] ==> finsum G (%i. f i ⊕ g i) {..n::nat} = finsum G f {..n} ⊕ finsum G g {..n}" by (rule comm_monoid.finprod_mult [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) lemma finsum_cong: "[| A = B; f : B -> carrier G; !!i. i : B =simp=> f i = g i |] ==> finsum G f A = finsum G g B" by (rule comm_monoid.finprod_cong [OF a_comm_monoid, folded finsum_def, simplified monoid_record_simps]) (auto simp add: simp_implies_def) text {*Usually, if this rule causes a failed congruence proof error, the reason is that the premise @{text "g ∈ B -> carrier G"} cannot be shown. Adding @{thm [source] Pi_def} to the simpset is often useful. *} lemma finsum_reindex: assumes fin: "finite A" shows "f : (h ` A) -> carrier G ==> inj_on h A ==> finsum G f (h ` A) = finsum G (%x. f (h x)) A" using fin apply induct apply (auto simp add: finsum_insert Pi_def) done (* The following is wrong. Needed is the equivalent of (^) for addition, or indeed the canonical embedding from Nat into the monoid. lemma finsum_const: assumes fin [simp]: "finite A" and a [simp]: "a : carrier G" shows "finsum G (%x. a) A = a (^) card A" using fin apply induct apply force apply (subst finsum_insert) apply auto apply (force simp add: Pi_def) apply (subst m_comm) apply auto done *) (* By Jesus Aransay. *) lemma finsum_singleton: assumes i_in_A: "i ∈ A" and fin_A: "finite A" and f_Pi: "f ∈ A -> carrier G" shows "(\<Oplus>j∈A. if i = j then f j else \<zero>) = f i" using i_in_A finsum_insert [of "A - {i}" i "(λj. if i = j then f j else \<zero>)"] fin_A f_Pi finsum_zero [of "A - {i}"] finsum_cong [of "A - {i}" "A - {i}" "(λj. if i = j then f j else \<zero>)" "(λi. \<zero>)"] unfolding Pi_def simp_implies_def by (force simp add: insert_absorb) end subsection {* Rings: Basic Definitions *} locale ring = abelian_group R + monoid R for R (structure) + assumes l_distr: "[| x ∈ carrier R; y ∈ carrier R; z ∈ carrier R |] ==> (x ⊕ y) ⊗ z = x ⊗ z ⊕ y ⊗ z" and r_distr: "[| x ∈ carrier R; y ∈ carrier R; z ∈ carrier R |] ==> z ⊗ (x ⊕ y) = z ⊗ x ⊕ z ⊗ y" locale cring = ring + comm_monoid R locale "domain" = cring + assumes one_not_zero [simp]: "\<one> ~= \<zero>" and integral: "[| a ⊗ b = \<zero>; a ∈ carrier R; b ∈ carrier R |] ==> a = \<zero> | b = \<zero>" locale field = "domain" + assumes field_Units: "Units R = carrier R - {\<zero>}" subsection {* Rings *} lemma ringI: fixes R (structure) assumes abelian_group: "abelian_group R" and monoid: "monoid R" and l_distr: "!!x y z. [| x ∈ carrier R; y ∈ carrier R; z ∈ carrier R |] ==> (x ⊕ y) ⊗ z = x ⊗ z ⊕ y ⊗ z" and r_distr: "!!x y z. [| x ∈ carrier R; y ∈ carrier R; z ∈ carrier R |] ==> z ⊗ (x ⊕ y) = z ⊗ x ⊕ z ⊗ y" shows "ring R" by (auto intro: ring.intro abelian_group.axioms ring_axioms.intro assms) lemma (in ring) is_abelian_group: "abelian_group R" .. lemma (in ring) is_monoid: "monoid R" by (auto intro!: monoidI m_assoc) lemma (in ring) is_ring: "ring R" by (rule ring_axioms) lemmas ring_record_simps = monoid_record_simps ring.simps lemma cringI: fixes R (structure) assumes abelian_group: "abelian_group R" and comm_monoid: "comm_monoid R" and l_distr: "!!x y z. [| x ∈ carrier R; y ∈ carrier R; z ∈ carrier R |] ==> (x ⊕ y) ⊗ z = x ⊗ z ⊕ y ⊗ z" shows "cring R" proof (intro cring.intro ring.intro) show "ring_axioms R" -- {* Right-distributivity follows from left-distributivity and commutativity. *} proof (rule ring_axioms.intro) fix x y z assume R: "x ∈ carrier R" "y ∈ carrier R" "z ∈ carrier R" note [simp] = comm_monoid.axioms [OF comm_monoid] abelian_group.axioms [OF abelian_group] abelian_monoid.a_closed from R have "z ⊗ (x ⊕ y) = (x ⊕ y) ⊗ z" by (simp add: comm_monoid.m_comm [OF comm_monoid.intro]) also from R have "... = x ⊗ z ⊕ y ⊗ z" by (simp add: l_distr) also from R have "... = z ⊗ x ⊕ z ⊗ y" by (simp add: comm_monoid.m_comm [OF comm_monoid.intro]) finally show "z ⊗ (x ⊕ y) = z ⊗ x ⊕ z ⊗ y" . qed (rule l_distr) qed (auto intro: cring.intro abelian_group.axioms comm_monoid.axioms ring_axioms.intro assms) (* lemma (in cring) is_comm_monoid: "comm_monoid R" by (auto intro!: comm_monoidI m_assoc m_comm) *) lemma (in cring) is_cring: "cring R" by (rule cring_axioms) subsubsection {* Normaliser for Rings *} lemma (in abelian_group) r_neg2: "[| x ∈ carrier G; y ∈ carrier G |] ==> x ⊕ (\<ominus> x ⊕ y) = y" proof - assume G: "x ∈ carrier G" "y ∈ carrier G" then have "(x ⊕ \<ominus> x) ⊕ y = y" by (simp only: r_neg l_zero) with G show ?thesis by (simp add: a_ac) qed lemma (in abelian_group) r_neg1: "[| x ∈ carrier G; y ∈ carrier G |] ==> \<ominus> x ⊕ (x ⊕ y) = y" proof - assume G: "x ∈ carrier G" "y ∈ carrier G" then have "(\<ominus> x ⊕ x) ⊕ y = y" by (simp only: l_neg l_zero) with G show ?thesis by (simp add: a_ac) qed text {* The following proofs are from Jacobson, Basic Algebra I, pp.~88--89 *} lemma (in ring) l_null [simp]: "x ∈ carrier R ==> \<zero> ⊗ x = \<zero>" proof - assume R: "x ∈ carrier R" then have "\<zero> ⊗ x ⊕ \<zero> ⊗ x = (\<zero> ⊕ \<zero>) ⊗ x" by (simp add: l_distr del: l_zero r_zero) also from R have "... = \<zero> ⊗ x ⊕ \<zero>" by simp finally have "\<zero> ⊗ x ⊕ \<zero> ⊗ x = \<zero> ⊗ x ⊕ \<zero>" . with R show ?thesis by (simp del: r_zero) qed lemma (in ring) r_null [simp]: "x ∈ carrier R ==> x ⊗ \<zero> = \<zero>" proof - assume R: "x ∈ carrier R" then have "x ⊗ \<zero> ⊕ x ⊗ \<zero> = x ⊗ (\<zero> ⊕ \<zero>)" by (simp add: r_distr del: l_zero r_zero) also from R have "... = x ⊗ \<zero> ⊕ \<zero>" by simp finally have "x ⊗ \<zero> ⊕ x ⊗ \<zero> = x ⊗ \<zero> ⊕ \<zero>" . with R show ?thesis by (simp del: r_zero) qed lemma (in ring) l_minus: "[| x ∈ carrier R; y ∈ carrier R |] ==> \<ominus> x ⊗ y = \<ominus> (x ⊗ y)" proof - assume R: "x ∈ carrier R" "y ∈ carrier R" then have "(\<ominus> x) ⊗ y ⊕ x ⊗ y = (\<ominus> x ⊕ x) ⊗ y" by (simp add: l_distr) also from R have "... = \<zero>" by (simp add: l_neg l_null) finally have "(\<ominus> x) ⊗ y ⊕ x ⊗ y = \<zero>" . with R have "(\<ominus> x) ⊗ y ⊕ x ⊗ y ⊕ \<ominus> (x ⊗ y) = \<zero> ⊕ \<ominus> (x ⊗ y)" by simp with R show ?thesis by (simp add: a_assoc r_neg) qed lemma (in ring) r_minus: "[| x ∈ carrier R; y ∈ carrier R |] ==> x ⊗ \<ominus> y = \<ominus> (x ⊗ y)" proof - assume R: "x ∈ carrier R" "y ∈ carrier R" then have "x ⊗ (\<ominus> y) ⊕ x ⊗ y = x ⊗ (\<ominus> y ⊕ y)" by (simp add: r_distr) also from R have "... = \<zero>" by (simp add: l_neg r_null) finally have "x ⊗ (\<ominus> y) ⊕ x ⊗ y = \<zero>" . with R have "x ⊗ (\<ominus> y) ⊕ x ⊗ y ⊕ \<ominus> (x ⊗ y) = \<zero> ⊕ \<ominus> (x ⊗ y)" by simp with R show ?thesis by (simp add: a_assoc r_neg ) qed lemma (in abelian_group) minus_eq: "[| x ∈ carrier G; y ∈ carrier G |] ==> x \<ominus> y = x ⊕ \<ominus> y" by (simp only: a_minus_def) text {* Setup algebra method: compute distributive normal form in locale contexts *} use "ringsimp.ML" setup Algebra.setup lemmas (in ring) ring_simprules [algebra ring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] = a_closed zero_closed a_inv_closed minus_closed m_closed one_closed a_assoc l_zero l_neg a_comm m_assoc l_one l_distr minus_eq r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero a_lcomm r_distr l_null r_null l_minus r_minus lemmas (in cring) [algebra del: ring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] = _ lemmas (in cring) cring_simprules [algebra add: cring "zero R" "add R" "a_inv R" "a_minus R" "one R" "mult R"] = a_closed zero_closed a_inv_closed minus_closed m_closed one_closed a_assoc l_zero l_neg a_comm m_assoc l_one l_distr m_comm minus_eq r_zero r_neg r_neg2 r_neg1 minus_add minus_minus minus_zero a_lcomm m_lcomm r_distr l_null r_null l_minus r_minus lemma (in cring) nat_pow_zero: "(n::nat) ~= 0 ==> \<zero> (^) n = \<zero>" by (induct n) simp_all lemma (in ring) one_zeroD: assumes onezero: "\<one> = \<zero>" shows "carrier R = {\<zero>}" proof (rule, rule) fix x assume xcarr: "x ∈ carrier R" from xcarr have "x = x ⊗ \<one>" by simp from this and onezero have "x = x ⊗ \<zero>" by simp from this and xcarr have "x = \<zero>" by simp thus "x ∈ {\<zero>}" by fast qed fast lemma (in ring) one_zeroI: assumes carrzero: "carrier R = {\<zero>}" shows "\<one> = \<zero>" proof - from one_closed and carrzero show "\<one> = \<zero>" by simp qed lemma (in ring) carrier_one_zero: shows "(carrier R = {\<zero>}) = (\<one> = \<zero>)" by (rule, erule one_zeroI, erule one_zeroD) lemma (in ring) carrier_one_not_zero: shows "(carrier R ≠ {\<zero>}) = (\<one> ≠ \<zero>)" by (simp add: carrier_one_zero) text {* Two examples for use of method algebra *} lemma fixes R (structure) and S (structure) assumes "ring R" "cring S" assumes RS: "a ∈ carrier R" "b ∈ carrier R" "c ∈ carrier S" "d ∈ carrier S" shows "a ⊕ \<ominus> (a ⊕ \<ominus> b) = b & c ⊗S d = d ⊗S c" proof - interpret ring R by fact interpret cring S by fact ML_val {* Algebra.print_structures @{context} *} from RS show ?thesis by algebra qed lemma fixes R (structure) assumes "ring R" assumes R: "a ∈ carrier R" "b ∈ carrier R" shows "a \<ominus> (a \<ominus> b) = b" proof - interpret ring R by fact from R show ?thesis by algebra qed subsubsection {* Sums over Finite Sets *} lemma (in ring) finsum_ldistr: "[| finite A; a ∈ carrier R; f ∈ A -> carrier R |] ==> finsum R f A ⊗ a = finsum R (%i. f i ⊗ a) A" proof (induct set: finite) case empty then show ?case by simp next case (insert x F) then show ?case by (simp add: Pi_def l_distr) qed lemma (in ring) finsum_rdistr: "[| finite A; a ∈ carrier R; f ∈ A -> carrier R |] ==> a ⊗ finsum R f A = finsum R (%i. a ⊗ f i) A" proof (induct set: finite) case empty then show ?case by simp next case (insert x F) then show ?case by (simp add: Pi_def r_distr) qed subsection {* Integral Domains *} lemma (in "domain") zero_not_one [simp]: "\<zero> ~= \<one>" by (rule not_sym) simp lemma (in "domain") integral_iff: (* not by default a simp rule! *) "[| a ∈ carrier R; b ∈ carrier R |] ==> (a ⊗ b = \<zero>) = (a = \<zero> | b = \<zero>)" proof assume "a ∈ carrier R" "b ∈ carrier R" "a ⊗ b = \<zero>" then show "a = \<zero> | b = \<zero>" by (simp add: integral) next assume "a ∈ carrier R" "b ∈ carrier R" "a = \<zero> | b = \<zero>" then show "a ⊗ b = \<zero>" by auto qed lemma (in "domain") m_lcancel: assumes prem: "a ~= \<zero>" and R: "a ∈ carrier R" "b ∈ carrier R" "c ∈ carrier R" shows "(a ⊗ b = a ⊗ c) = (b = c)" proof assume eq: "a ⊗ b = a ⊗ c" with R have "a ⊗ (b \<ominus> c) = \<zero>" by algebra with R have "a = \<zero> | (b \<ominus> c) = \<zero>" by (simp add: integral_iff) with prem and R have "b \<ominus> c = \<zero>" by auto with R have "b = b \<ominus> (b \<ominus> c)" by algebra also from R have "b \<ominus> (b \<ominus> c) = c" by algebra finally show "b = c" . next assume "b = c" then show "a ⊗ b = a ⊗ c" by simp qed lemma (in "domain") m_rcancel: assumes prem: "a ~= \<zero>" and R: "a ∈ carrier R" "b ∈ carrier R" "c ∈ carrier R" shows conc: "(b ⊗ a = c ⊗ a) = (b = c)" proof - from prem and R have "(a ⊗ b = a ⊗ c) = (b = c)" by (rule m_lcancel) with R show ?thesis by algebra qed subsection {* Fields *} text {* Field would not need to be derived from domain, the properties for domain follow from the assumptions of field *} lemma (in cring) cring_fieldI: assumes field_Units: "Units R = carrier R - {\<zero>}" shows "field R" proof from field_Units have a: "\<zero> ∉ Units R" by fast have "\<one> ∈ Units R" by fast from this and a show "\<one> ≠ \<zero>" by force next fix a b assume acarr: "a ∈ carrier R" and bcarr: "b ∈ carrier R" and ab: "a ⊗ b = \<zero>" show "a = \<zero> ∨ b = \<zero>" proof (cases "a = \<zero>", simp) assume "a ≠ \<zero>" from this and field_Units and acarr have aUnit: "a ∈ Units R" by fast from bcarr have "b = \<one> ⊗ b" by algebra also from aUnit acarr have "... = (inv a ⊗ a) ⊗ b" by (simp add: Units_l_inv) also from acarr bcarr aUnit[THEN Units_inv_closed] have "... = (inv a) ⊗ (a ⊗ b)" by algebra also from ab and acarr bcarr aUnit have "... = (inv a) ⊗ \<zero>" by simp also from aUnit[THEN Units_inv_closed] have "... = \<zero>" by algebra finally have "b = \<zero>" . thus "a = \<zero> ∨ b = \<zero>" by simp qed qed (rule field_Units) text {* Another variant to show that something is a field *} lemma (in cring) cring_fieldI2: assumes notzero: "\<zero> ≠ \<one>" and invex: "!!a. [|a ∈ carrier R; a ≠ \<zero>|] ==> ∃b∈carrier R. a ⊗ b = \<one>" shows "field R" apply (rule cring_fieldI, simp add: Units_def) apply (rule, clarsimp) apply (simp add: notzero) proof (clarsimp) fix x assume xcarr: "x ∈ carrier R" and "x ≠ \<zero>" from this have "∃y∈carrier R. x ⊗ y = \<one>" by (rule invex) from this obtain y where ycarr: "y ∈ carrier R" and xy: "x ⊗ y = \<one>" by fast from xy xcarr ycarr have "y ⊗ x = \<one>" by (simp add: m_comm) from ycarr and this and xy show "∃y∈carrier R. y ⊗ x = \<one> ∧ x ⊗ y = \<one>" by fast qed subsection {* Morphisms *} constdefs (structure R S) ring_hom :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme] => ('a => 'b) set" "ring_hom R S == {h. h ∈ carrier R -> carrier S & (ALL x y. x ∈ carrier R & y ∈ carrier R --> h (x ⊗ y) = h x ⊗S h y & h (x ⊕ y) = h x ⊕S h y) & h \<one> = \<one>S}" lemma ring_hom_memI: fixes R (structure) and S (structure) assumes hom_closed: "!!x. x ∈ carrier R ==> h x ∈ carrier S" and hom_mult: "!!x y. [| x ∈ carrier R; y ∈ carrier R |] ==> h (x ⊗ y) = h x ⊗S h y" and hom_add: "!!x y. [| x ∈ carrier R; y ∈ carrier R |] ==> h (x ⊕ y) = h x ⊕S h y" and hom_one: "h \<one> = \<one>S" shows "h ∈ ring_hom R S" by (auto simp add: ring_hom_def assms Pi_def) lemma ring_hom_closed: "[| h ∈ ring_hom R S; x ∈ carrier R |] ==> h x ∈ carrier S" by (auto simp add: ring_hom_def funcset_mem) lemma ring_hom_mult: fixes R (structure) and S (structure) shows "[| h ∈ ring_hom R S; x ∈ carrier R; y ∈ carrier R |] ==> h (x ⊗ y) = h x ⊗S h y" by (simp add: ring_hom_def) lemma ring_hom_add: fixes R (structure) and S (structure) shows "[| h ∈ ring_hom R S; x ∈ carrier R; y ∈ carrier R |] ==> h (x ⊕ y) = h x ⊕S h y" by (simp add: ring_hom_def) lemma ring_hom_one: fixes R (structure) and S (structure) shows "h ∈ ring_hom R S ==> h \<one> = \<one>S" by (simp add: ring_hom_def) locale ring_hom_cring = R: cring R + S: cring S for R (structure) and S (structure) + fixes h assumes homh [simp, intro]: "h ∈ ring_hom R S" notes hom_closed [simp, intro] = ring_hom_closed [OF homh] and hom_mult [simp] = ring_hom_mult [OF homh] and hom_add [simp] = ring_hom_add [OF homh] and hom_one [simp] = ring_hom_one [OF homh] lemma (in ring_hom_cring) hom_zero [simp]: "h \<zero> = \<zero>S" proof - have "h \<zero> ⊕S h \<zero> = h \<zero> ⊕S \<zero>S" by (simp add: hom_add [symmetric] del: hom_add) then show ?thesis by (simp del: S.r_zero) qed lemma (in ring_hom_cring) hom_a_inv [simp]: "x ∈ carrier R ==> h (\<ominus> x) = \<ominus>S h x" proof - assume R: "x ∈ carrier R" then have "h x ⊕S h (\<ominus> x) = h x ⊕S (\<ominus>S h x)" by (simp add: hom_add [symmetric] R.r_neg S.r_neg del: hom_add) with R show ?thesis by simp qed lemma (in ring_hom_cring) hom_finsum [simp]: "[| finite A; f ∈ A -> carrier R |] ==> h (finsum R f A) = finsum S (h o f) A" proof (induct set: finite) case empty then show ?case by simp next case insert then show ?case by (simp add: Pi_def) qed lemma (in ring_hom_cring) hom_finprod: "[| finite A; f ∈ A -> carrier R |] ==> h (finprod R f A) = finprod S (h o f) A" proof (induct set: finite) case empty then show ?case by simp next case insert then show ?case by (simp add: Pi_def) qed declare ring_hom_cring.hom_finprod [simp] lemma id_ring_hom [simp]: "id ∈ ring_hom R R" by (auto intro!: ring_hom_memI) end