| author | desharna | 
| Fri, 21 Mar 2025 15:20:13 +0100 | |
| changeset 82314 | c95eca07f6a0 | 
| parent 81438 | 95c9af7483b1 | 
| permissions | -rw-r--r-- | 
| 68582 | 1 | (* Title: HOL/Algebra/Chinese_Remainder.thy | 
| 2 | Author: Paulo Emílio de Vilhena | |
| 3 | *) | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | theory Chinese_Remainder | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 6 | imports Weak_Morphisms Ideal_Product | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 7 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | begin | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 11 | section \<open>Direct Product of Rings\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 12 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 13 | subsection \<open>Definitions\<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 15 | definition RDirProd :: "('a, 'n) ring_scheme \<Rightarrow> ('b, 'm) ring_scheme \<Rightarrow> ('a \<times> 'b) ring"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 16 | where "RDirProd R S = monoid.extend (R \<times>\<times> S) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 17 | \<lparr> zero = one ((add_monoid R) \<times>\<times> (add_monoid S)), | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 18 | add = mult ((add_monoid R) \<times>\<times> (add_monoid S)) \<rparr> " | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 19 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 20 | abbreviation nil_ring :: "('a list) ring"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 21 | where "nil_ring \<equiv> monoid.extend nil_monoid \<lparr> zero = [], add = (\<lambda>a b. []) \<rparr>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 22 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 23 | definition RDirProd_list :: "(('a, 'n) ring_scheme) list \<Rightarrow> ('a list) ring"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 24 | where "RDirProd_list Rs = foldr (\<lambda>R S. image_ring (\<lambda>(a, as). a # as) (RDirProd R S)) Rs nil_ring" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 25 | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 27 | subsection \<open>Basic Properties\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 28 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 29 | lemma RDirProd_carrier: "carrier (RDirProd R S) = carrier R \<times> carrier S" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 30 | unfolding RDirProd_def DirProd_def by (simp add: monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 31 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 32 | lemma RDirProd_add_monoid [simp]: "add_monoid (RDirProd R S) = (add_monoid R) \<times>\<times> (add_monoid S)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 33 | by (simp add: RDirProd_def monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 35 | lemma RDirProd_ring: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 36 | assumes "ring R" and "ring S" shows "ring (RDirProd R S)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 38 | have "monoid (RDirProd R S)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 39 | using DirProd_monoid[OF assms[THEN ring.axioms(2)]] unfolding monoid_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 40 | by (auto simp add: DirProd_def RDirProd_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 41 | then interpret Prod: group "add_monoid (RDirProd R S)" + monoid "RDirProd R S" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 42 | using DirProd_group[OF assms[THEN abelian_group.a_group[OF ring.is_abelian_group]]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 43 | unfolding RDirProd_add_monoid by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | show ?thesis | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 45 | by (unfold_locales, auto simp add: RDirProd_def DirProd_def monoid.defs assms ring.ring_simprules) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 48 | lemma RDirProd_iso1: | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | "(\<lambda>(x, y). (y, x)) \<in> ring_iso (RDirProd R S) (RDirProd S R)" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 50 | unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 51 | by (auto simp add: RDirProd_def DirProd_def monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 53 | lemma RDirProd_iso2: | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | "(\<lambda>(x, (y, z)). ((x, y), z)) \<in> ring_iso (RDirProd R (RDirProd S T)) (RDirProd (RDirProd R S) T)" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 55 | unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 56 | by (auto simp add: image_iff RDirProd_def DirProd_def monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 58 | lemma RDirProd_iso3: | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | "(\<lambda>((x, y), z). (x, (y, z))) \<in> ring_iso (RDirProd (RDirProd R S) T) (RDirProd R (RDirProd S T))" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 60 | unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 61 | by (auto simp add: image_iff RDirProd_def DirProd_def monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 63 | lemma RDirProd_iso4: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 64 | assumes "f \<in> ring_iso R S" shows "(\<lambda>(r, t). (f r, t)) \<in> ring_iso (RDirProd R T) (RDirProd S T)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 65 | using assms unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 66 | by (auto simp add: image_iff RDirProd_def DirProd_def monoid.defs)+ | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 68 | lemma RDirProd_iso5: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 69 | assumes "f \<in> ring_iso S T" shows "(\<lambda>(r, s). (r, f s)) \<in> ring_iso (RDirProd R S) (RDirProd R T)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 70 | using ring_iso_set_trans[OF ring_iso_set_trans[OF RDirProd_iso1 RDirProd_iso4[OF assms]] RDirProd_iso1] | 
| 68608 
4a4c2bc4b869
final removal of smt from Algebra
 paulson <lp15@cam.ac.uk> parents: 
68606diff
changeset | 71 | by (simp add: case_prod_unfold comp_def) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 73 | lemma RDirProd_iso6: | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | assumes "f \<in> ring_iso R R'" and "g \<in> ring_iso S S'" | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | shows "(\<lambda>(r, s). (f r, g s)) \<in> ring_iso (RDirProd R S) (RDirProd R' S')" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 76 | using ring_iso_set_trans[OF RDirProd_iso4[OF assms(1)] RDirProd_iso5[OF assms(2)]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 77 | by (simp add: case_prod_beta' comp_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 78 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 79 | lemma RDirProd_iso7: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 80 | shows "(\<lambda>a. (a, [])) \<in> ring_iso R (RDirProd R nil_ring)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 81 | unfolding ring_iso_def ring_hom_def bij_betw_def inj_on_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 82 | by (auto simp add: RDirProd_def DirProd_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 83 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 84 | lemma RDirProd_hom1: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 85 | shows "(\<lambda>a. (a, a)) \<in> ring_hom R (RDirProd R R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 86 | by (auto simp add: ring_hom_def RDirProd_def DirProd_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 87 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 88 | lemma RDirProd_hom2: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 89 | assumes "f \<in> ring_hom S T" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 90 | shows "(\<lambda>(x, y). (x, f y)) \<in> ring_hom (RDirProd R S) (RDirProd R T)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 91 | and "(\<lambda>(x, y). (f x, y)) \<in> ring_hom (RDirProd S R) (RDirProd T R)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 92 | using assms by (auto simp add: ring_hom_def RDirProd_def DirProd_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 93 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 94 | lemma RDirProd_hom3: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 95 | assumes "f \<in> ring_hom R R'" and "g \<in> ring_hom S S'" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 96 | shows "(\<lambda>(r, s). (f r, g s)) \<in> ring_hom (RDirProd R S) (RDirProd R' S')" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 97 | using ring_hom_trans[OF RDirProd_hom2(2)[OF assms(1)] RDirProd_hom2(1)[OF assms(2)]] | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | by (simp add: case_prod_beta' comp_def) | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 101 | subsection \<open>Direct Product of a List of Rings\<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 103 | lemma RDirProd_list_nil [simp]: "RDirProd_list [] = nil_ring" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 104 | unfolding RDirProd_list_def by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 106 | lemma nil_ring_simprules [simp]: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 107 |   "carrier nil_ring = { [] }" and "one nil_ring = []" and "zero nil_ring = []"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 108 | by (auto simp add: monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 110 | lemma RDirProd_list_truncate: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 111 | shows "monoid.truncate (RDirProd_list Rs) = DirProd_list Rs" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 112 | proof (induct Rs, simp add: RDirProd_list_def DirProd_list_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 113 | case (Cons R Rs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 114 | have "monoid.truncate (RDirProd_list (R # Rs)) = | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 115 | monoid.truncate (image_ring (\<lambda>(a, as). a # as) (RDirProd R (RDirProd_list Rs)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 116 | unfolding RDirProd_list_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 117 | also have " ... = image_group (\<lambda>(a, as). a # as) (monoid.truncate (RDirProd R (RDirProd_list Rs)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 118 | by (simp add: image_ring_def image_group_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 119 | also have " ... = image_group (\<lambda>(a, as). a # as) (R \<times>\<times> (monoid.truncate (RDirProd_list Rs)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 120 | by (simp add: RDirProd_def DirProd_def monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 121 | also have " ... = DirProd_list (R # Rs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 122 | unfolding Cons DirProd_list_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 123 | finally show ?case . | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 126 | lemma RDirProd_list_carrier_def': | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 127 | shows "carrier (RDirProd_list Rs) = carrier (DirProd_list Rs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 128 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 129 | have "carrier (RDirProd_list Rs) = carrier (monoid.truncate (RDirProd_list Rs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 130 | by (simp add: monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 131 | thus ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 132 | unfolding RDirProd_list_truncate . | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 135 | lemma RDirProd_list_carrier: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 136 | shows "carrier (RDirProd_list (G # Gs)) = (\<lambda>(x, xs). x # xs) ` (carrier G \<times> carrier (RDirProd_list Gs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 137 | unfolding RDirProd_list_carrier_def' using DirProd_list_carrier . | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 138 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 139 | lemma RDirProd_list_one: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 140 | shows "one (RDirProd_list Rs) = foldr (\<lambda>R tl. (one R) # tl) Rs []" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 141 | unfolding RDirProd_list_def RDirProd_def image_ring_def image_group_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 142 | by (induct Rs) (auto simp add: monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 144 | lemma RDirProd_list_zero: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 145 | shows "zero (RDirProd_list Rs) = foldr (\<lambda>R tl. (zero R) # tl) Rs []" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 146 | unfolding RDirProd_list_def RDirProd_def image_ring_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 147 | by (induct Rs) (auto simp add: monoid.defs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 148 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 149 | lemma RDirProd_list_zero': | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 150 | shows "zero (RDirProd_list (R # Rs)) = (zero R) # (zero (RDirProd_list Rs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 151 | unfolding RDirProd_list_zero by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 152 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 153 | lemma RDirProd_list_carrier_mem: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 154 | assumes "as \<in> carrier (RDirProd_list Rs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 155 | shows "length as = length Rs" and "\<And>i. i < length Rs \<Longrightarrow> (as ! i) \<in> carrier (Rs ! i)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 156 | using assms DirProd_list_carrier_mem unfolding RDirProd_list_carrier_def' by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 158 | lemma RDirProd_list_carrier_memI: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 159 | assumes "length as = length Rs" and "\<And>i. i < length Rs \<Longrightarrow> (as ! i) \<in> carrier (Rs ! i)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 160 | shows "as \<in> carrier (RDirProd_list Rs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 161 | using assms DirProd_list_carrier_memI unfolding RDirProd_list_carrier_def' by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 162 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 163 | lemma inj_on_RDirProd_carrier: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 164 | shows "inj_on (\<lambda>(a, as). a # as) (carrier (RDirProd R (RDirProd_list Rs)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 165 | unfolding RDirProd_def DirProd_def inj_on_def by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 167 | lemma RDirProd_list_is_ring: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 168 | assumes "\<And>i. i < length Rs \<Longrightarrow> ring (Rs ! i)" shows "ring (RDirProd_list Rs)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 169 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 170 | proof (induct Rs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 171 | case Nil thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 172 | unfolding RDirProd_list_def by (unfold_locales, auto simp add: monoid.defs Units_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 173 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 174 | case (Cons R Rs) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 175 | hence is_ring: "ring (RDirProd R (RDirProd_list Rs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 176 | using RDirProd_ring[of R "RDirProd_list Rs"] by force | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 177 | show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 178 | using ring.inj_imp_image_ring_is_ring[OF is_ring inj_on_RDirProd_carrier] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 179 | unfolding RDirProd_list_def by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 182 | lemma RDirProd_list_iso1: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 183 | "(\<lambda>(a, as). a # as) \<in> ring_iso (RDirProd R (RDirProd_list Rs)) (RDirProd_list (R # Rs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 184 | using inj_imp_image_ring_iso[OF inj_on_RDirProd_carrier] unfolding RDirProd_list_def by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 186 | lemma RDirProd_list_iso2: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 187 | "Hilbert_Choice.inv (\<lambda>(a, as). a # as) \<in> ring_iso (RDirProd_list (R # Rs)) (RDirProd R (RDirProd_list Rs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 188 | unfolding RDirProd_list_def by (auto intro: inj_imp_image_ring_inv_iso simp add: inj_def) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 190 | lemma RDirProd_list_iso3: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 191 | "(\<lambda>a. [ a ]) \<in> ring_iso R (RDirProd_list [ R ])" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 192 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 193 | have [simp]: "(\<lambda>a. [ a ]) = (\<lambda>(a, as). a # as) \<circ> (\<lambda>a. (a, []))" by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 194 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 195 | using ring_iso_set_trans[OF RDirProd_iso7] RDirProd_list_iso1[of R "[]"] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 196 | unfolding RDirProd_list_def by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 197 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 198 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 199 | lemma RDirProd_list_hom1: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 200 | "(\<lambda>(a, as). a # as) \<in> ring_hom (RDirProd R (RDirProd_list Rs)) (RDirProd_list (R # Rs))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 201 | using RDirProd_list_iso1 unfolding ring_iso_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 202 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 203 | lemma RDirProd_list_hom2: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 204 | assumes "f \<in> ring_hom R S" shows "(\<lambda>a. [ f a ]) \<in> ring_hom R (RDirProd_list [ S ])" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 205 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 206 | have hom1: "(\<lambda>a. (a, [])) \<in> ring_hom R (RDirProd R nil_ring)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 207 | using RDirProd_iso7 unfolding ring_iso_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 208 | have hom2: "(\<lambda>(a, as). a # as) \<in> ring_hom (RDirProd S nil_ring) (RDirProd_list [ S ])" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 209 | using RDirProd_list_hom1[of _ "[]"] unfolding RDirProd_list_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 210 | have [simp]: "(\<lambda>(a, as). a # as) \<circ> ((\<lambda>(x, y). (f x, y)) \<circ> (\<lambda>a. (a, []))) = (\<lambda>a. [ f a ])" by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 211 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 212 | using ring_hom_trans[OF ring_hom_trans[OF hom1 RDirProd_hom2(2)[OF assms]] hom2] by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 216 | section \<open>Chinese Remainder Theorem\<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 218 | subsection \<open>Definitions\<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 220 | abbreviation (in ring) canonical_proj :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'a set \<times> 'a set" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 221 | where "canonical_proj I J \<equiv> (\<lambda>a. (I +> a, J +> a))" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 223 | definition (in ring) canonical_proj_ext :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> ('a set) list"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 224 | where "canonical_proj_ext I n = (\<lambda>a. map (\<lambda>i. (I i) +> a) [0..< Suc n])" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 227 | subsection \<open>Chinese Remainder Simple\<close> | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 229 | lemma (in ring) canonical_proj_is_surj: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 230 | assumes "ideal I R" "ideal J R" and "I <+> J = carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 231 | shows "(canonical_proj I J) ` carrier R = carrier (RDirProd (R Quot I) (R Quot J))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 232 | unfolding RDirProd_def DirProd_def FactRing_def A_RCOSETS_def' | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 233 | proof (auto simp add: monoid.defs) | 
| 81438 | 234 | have aux_lemma1: "I +> i = \<zero>\<^bsub>R Quot I\<^esub>" if "ideal I R" "i \<in> I" for I i | 
| 235 | using that a_rcos_zero by (simp add: FactRing_def) | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | |
| 81438 | 237 | have aux_lemma2: "I +> j = \<one>\<^bsub>R Quot I\<^esub>" | 
| 238 | if A: "ideal I R" "i \<in> I" "j \<in> carrier R" "i \<oplus> j = \<one>" | |
| 239 | for I i j | |
| 240 | proof - | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 241 | have "(I +> i) \<oplus>\<^bsub>R Quot I\<^esub> (I +> j) = I +> \<one>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 242 | using ring_hom_memE(3)[OF ideal.rcos_ring_hom ideal.Icarr[OF _ A(2)] A(3)] A(1,4) by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 243 | moreover have "I +> i = I" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 244 | using abelian_subgroupI3[OF ideal.axioms(1) is_abelian_group] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 245 | by (simp add: A(1-2) abelian_subgroup.a_rcos_const) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 246 | moreover have "I +> j \<in> carrier (R Quot I)" and "I = \<zero>\<^bsub>R Quot I\<^esub>" and "I +> \<one> = \<one>\<^bsub>R Quot I\<^esub>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 247 | by (auto simp add: FactRing_def A_RCOSETS_def' A(3)) | 
| 81438 | 248 | ultimately show ?thesis | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 249 | using ring.ring_simprules(8)[OF ideal.quotient_is_ring[OF A(1)]] by simp | 
| 81438 | 250 | qed | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 252 | interpret I: ring "R Quot I" + J: ring "R Quot J" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 253 | using assms(1-2)[THEN ideal.quotient_is_ring] by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 255 | fix a b assume a: "a \<in> carrier R" and b: "b \<in> carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 256 | have "\<one> \<in> I <+> J" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 257 | using assms(3) by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 258 | then obtain i j where i: "i \<in> carrier R" "i \<in> I" and j: "j \<in> carrier R" "j \<in> J" and ij: "i \<oplus> j = \<one>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 259 | using assms(1-2)[THEN ideal.Icarr] unfolding set_add_def' by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 260 | hence rcos_j: "I +> j = \<one>\<^bsub>R Quot I\<^esub>" and rcos_i: "J +> i = \<one>\<^bsub>R Quot J\<^esub>" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 261 | using assms(1-2)[THEN aux_lemma2] a_comm by simp+ | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 263 | define s where "s = (a \<otimes> j) \<oplus> (b \<otimes> i)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 264 | hence "s \<in> carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 265 | using a b i j by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 267 | have "I +> s = ((I +> a) \<otimes>\<^bsub>R Quot I\<^esub> (I +> j)) \<oplus>\<^bsub>R Quot I\<^esub> (I +> (b \<otimes> i))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 268 | using ring_hom_memE(2-3)[OF ideal.rcos_ring_hom[OF assms(1)]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 269 | by (simp add: a b i(1) j(1) s_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 270 | moreover have "I +> a \<in> carrier (R Quot I)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 271 | by (auto simp add: FactRing_def A_RCOSETS_def' a) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 272 | ultimately have "I +> s = I +> a" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 273 | unfolding rcos_j aux_lemma1[OF assms(1) ideal.I_l_closed[OF assms(1) i(2) b]] by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 275 | have "J +> s = (J +> (a \<otimes> j)) \<oplus>\<^bsub>R Quot J\<^esub> ((J +> b) \<otimes>\<^bsub>R Quot J\<^esub> (J +> i))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 276 | using ring_hom_memE(2-3)[OF ideal.rcos_ring_hom[OF assms(2)]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 277 | by (simp add: a b i(1) j(1) s_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 278 | moreover have "J +> b \<in> carrier (R Quot J)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 279 | by (auto simp add: FactRing_def A_RCOSETS_def' b) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 280 | ultimately have "J +> s = J +> b" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 281 | unfolding rcos_i aux_lemma1[OF assms(2) ideal.I_l_closed[OF assms(2) j(2) a]] by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 283 | from \<open>I +> s = I +> a\<close> and \<open>J +> s = J +> b\<close> and \<open>s \<in> carrier R\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 284 | show "(I +> a, J +> b) \<in> (canonical_proj I J) ` carrier R" by blast | 
| 68606 
96a49db47c97
removal of smt and certain refinements
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 285 | qed | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 287 | lemma (in ring) canonical_proj_ker: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 288 | assumes "ideal I R" and "ideal J R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 289 | shows "a_kernel R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J) = I \<inter> J" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 290 | proof | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 291 | show "a_kernel R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J) \<subseteq> I \<inter> J" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 292 | unfolding FactRing_def RDirProd_def DirProd_def a_kernel_def' | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 293 | by (auto simp add: assms[THEN ideal.rcos_const_imp_mem] monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 295 | show "I \<inter> J \<subseteq> a_kernel R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 296 | proof | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 297 | fix s assume s: "s \<in> I \<inter> J" then have "I +> s = I" and "J +> s = J" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 298 | using abelian_subgroupI3[OF ideal.axioms(1) is_abelian_group] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 299 | by (simp add: abelian_subgroup.a_rcos_const assms)+ | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 300 | thus "s \<in> a_kernel R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 301 | unfolding FactRing_def RDirProd_def DirProd_def a_kernel_def' | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 302 | using ideal.Icarr[OF assms(1)] s by (simp add: monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 306 | lemma (in ring) canonical_proj_is_hom: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 307 | assumes "ideal I R" and "ideal J R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 308 | shows "(canonical_proj I J) \<in> ring_hom R (RDirProd (R Quot I) (R Quot J))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 309 | unfolding RDirProd_def DirProd_def FactRing_def A_RCOSETS_def' | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 310 | by (auto intro!: ring_hom_memI | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 311 | simp add: assms[THEN ideal.rcoset_mult_add] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 312 | assms[THEN ideal.a_rcos_sum] monoid.defs) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 314 | lemma (in ring) canonical_proj_ring_hom: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 315 | assumes "ideal I R" and "ideal J R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 316 | shows "ring_hom_ring R (RDirProd (R Quot I) (R Quot J)) (canonical_proj I J)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 317 | using ring_hom_ring.intro[OF ring_axioms RDirProd_ring[OF assms[THEN ideal.quotient_is_ring]]] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 318 | by (simp add: ring_hom_ring_axioms_def canonical_proj_is_hom[OF assms]) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 320 | theorem (in ring) chinese_remainder_simple: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 321 | assumes "ideal I R" "ideal J R" and "I <+> J = carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 322 | shows "R Quot (I \<inter> J) \<simeq> RDirProd (R Quot I) (R Quot J)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 323 | using ring_hom_ring.FactRing_iso[OF canonical_proj_ring_hom canonical_proj_is_surj] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 324 | by (simp add: canonical_proj_ker assms) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 326 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 327 | subsection \<open>Chinese Remainder Generalized\<close> | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 328 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 329 | lemma (in ring) canonical_proj_ext_zero [simp]: "(canonical_proj_ext I 0) = (\<lambda>a. [ (I 0) +> a ])" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 330 | unfolding canonical_proj_ext_def by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 332 | lemma (in ring) canonical_proj_ext_tl: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 333 | "(\<lambda>a. canonical_proj_ext I (Suc n) a) = (\<lambda>a. ((I 0) +> a) # (canonical_proj_ext (\<lambda>i. I (Suc i)) n a))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 334 | unfolding canonical_proj_ext_def by (induct n) (auto, metis (lifting) append.assoc append_Cons append_Nil) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 335 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 336 | lemma (in ring) canonical_proj_ext_is_hom: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 337 | assumes "\<And>i. i \<le> n \<Longrightarrow> ideal (I i) R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 338 | shows "(canonical_proj_ext I n) \<in> ring_hom R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 339 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 340 | proof (induct n arbitrary: I) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 341 | case 0 thus ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 342 | using RDirProd_list_hom2[OF ideal.rcos_ring_hom[of _ R]] by (simp add: canonical_proj_ext_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 343 | next | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 344 | let ?DirProd = "\<lambda>I n. RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..<Suc n])" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 345 | let ?proj = "\<lambda>I n. canonical_proj_ext I n" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 347 | case (Suc n) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 348 | hence I: "ideal (I 0) R" by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 349 | have hom: "(?proj (\<lambda>i. I (Suc i)) n) \<in> ring_hom R (?DirProd (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 350 | using Suc(1)[of "\<lambda>i. I (Suc i)"] Suc(2) by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 351 | have [simp]: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 352 | "(\<lambda>(a, as). a # as) \<circ> ((\<lambda>(r, s). (I 0 +> r, ?proj (\<lambda>i. I (Suc i)) n s)) \<circ> (\<lambda>a. (a, a))) = ?proj I (Suc n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 353 | unfolding canonical_proj_ext_tl by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 354 | moreover have | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 355 | "(R Quot I 0) # (map (\<lambda>i. R Quot I (Suc i)) [0..< Suc n]) = map (\<lambda>i. R Quot (I i)) [0..< Suc (Suc n)]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 356 | by (induct n) (auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 357 | moreover show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 358 | using ring_hom_trans[OF ring_hom_trans[OF RDirProd_hom1 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 359 | RDirProd_hom3[OF ideal.rcos_ring_hom[OF I] hom]] RDirProd_list_hom1] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 360 | unfolding calculation(2) by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 363 | lemma (in ring) RDirProd_Quot_list_is_ring: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 364 | assumes "\<And>i. i \<le> n \<Longrightarrow> ideal (I i) R" shows "ring (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 365 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 366 | have ring_list: "\<And>i. i < Suc n \<Longrightarrow> ring ((map (\<lambda>i. R Quot I i) [0..< Suc n]) ! i)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 367 | using ideal.quotient_is_ring[OF assms] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 368 | by (metis add.left_neutral diff_zero le_simps(2) nth_map_upt) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 369 | show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 370 | using RDirProd_list_is_ring[OF ring_list] by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 371 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 372 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 373 | lemma (in ring) canonical_proj_ext_ring_hom: | 
| 68606 
96a49db47c97
removal of smt and certain refinements
 paulson <lp15@cam.ac.uk> parents: 
68582diff
changeset | 374 | assumes "\<And>i. i \<le> n \<Longrightarrow> ideal (I i) R" | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 375 | shows "ring_hom_ring R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])) (canonical_proj_ext I n)" | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | proof - | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 377 | have ring: "ring (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 378 | using RDirProd_Quot_list_is_ring[OF assms] by simp | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | show ?thesis | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 380 | using canonical_proj_ext_is_hom assms ring_hom_ring.intro[OF ring_axioms ring] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 381 | unfolding ring_hom_ring_axioms_def by blast | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 382 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 383 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 384 | lemma (in ring) canonical_proj_ext_ker: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 385 | assumes "\<And>i. i \<le> (n :: nat) \<Longrightarrow> ideal (I i) R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 386 | shows "a_kernel R (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])) (canonical_proj_ext I n) = (\<Inter>i \<le> n. I i)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 387 | proof - | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 388 | let ?map_Quot = "\<lambda>I n. map (\<lambda>i. R Quot (I i)) [0..< Suc n]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 389 | let ?ker = "\<lambda>I n. a_kernel R (RDirProd_list (?map_Quot I n)) (canonical_proj_ext I n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 390 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 391 | from assms show ?thesis | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 392 | proof (induct n arbitrary: I) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 393 | case 0 then have I: "ideal (I 0) R" by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 394 | show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 395 | unfolding a_kernel_def' RDirProd_list_zero canonical_proj_ext_def FactRing_def | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 396 | using ideal.rcos_const_imp_mem a_rcos_zero ideal.Icarr I by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | next | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 398 | case (Suc n) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 399 | hence I: "ideal (I 0) R" by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 400 | have map_simp: "?map_Quot I (Suc n) = (R Quot I 0) # (?map_Quot (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 401 | by (induct n) (auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 402 | have ker_I0: "I 0 = a_kernel R (R Quot (I 0)) (\<lambda>a. (I 0) +> a)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 403 | using ideal.rcos_const_imp_mem[OF I] a_rcos_zero[OF I] ideal.Icarr[OF I] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 404 | unfolding a_kernel_def' FactRing_def by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 405 | hence "?ker I (Suc n) = (?ker (\<lambda>i. I (Suc i)) n) \<inter> (I 0)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 406 | unfolding a_kernel_def' map_simp RDirProd_list_zero' canonical_proj_ext_tl by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 407 | moreover have "?ker (\<lambda>i. I (Suc i)) n = (\<Inter>i \<le> n. I (Suc i))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 408 | using Suc(1)[of "\<lambda>i. I (Suc i)"] Suc(2) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 409 | ultimately show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 410 | by (auto simp add: INT_extend_simps(10) atMost_atLeast0) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 411 | (metis atLeastAtMost_iff le_zero_eq not_less_eq_eq) | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 413 | qed | 
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 415 | lemma (in cring) canonical_proj_ext_is_surj: | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 416 | assumes "\<And>i. i \<le> n \<Longrightarrow> ideal (I i) R" and "\<And>i j. \<lbrakk> i \<le> n; j \<le> n \<rbrakk> \<Longrightarrow> i \<noteq> j \<Longrightarrow> I i <+> I j = carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 417 | shows "(canonical_proj_ext I n) ` carrier R = carrier (RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n]))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 418 | using assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 419 | proof (induct n arbitrary: I) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 420 | case 0 show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 421 | by (auto simp add: RDirProd_list_carrier FactRing_def A_RCOSETS_def') | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 422 | next | 
| 81438 | 423 | have aux_lemma: "(\<lambda>a. (f a, g a)) ` carrier R = (f ` carrier R) \<times> (g ` carrier R)" | 
| 424 | if A: "ring T" "f \<in> ring_hom R S" "g \<in> ring_hom R T" "f ` carrier R \<subseteq> f ` (a_kernel R T g)" | |
| 425 | for S :: "'c ring" and T :: "'d ring" and f g | |
| 426 | proof | |
| 427 | show "(\<lambda>a. (f a, g a)) ` carrier R \<subseteq> (f ` carrier R) \<times> (g ` carrier R)" | |
| 428 | by blast | |
| 429 | next | |
| 430 | show "(f ` carrier R) \<times> (g ` carrier R) \<subseteq> (\<lambda>a. (f a, g a)) ` carrier R" | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 431 | proof | 
| 81438 | 432 | fix t assume "t \<in> (f ` carrier R) \<times> (g ` carrier R)" | 
| 433 | then obtain a b where a: "a \<in> carrier R" "f a = fst t" and b: "b \<in> carrier R" "g b = snd t" | |
| 434 | by auto | |
| 435 | obtain c where c: "c \<in> a_kernel R T g" "f c = f (a \<ominus> b)" | |
| 436 | using A(4) minus_closed[OF a(1) b (1)] by auto | |
| 437 | have "f (c \<oplus> b) = f (a \<ominus> b) \<oplus>\<^bsub>S\<^esub> f b" | |
| 438 | using ring_hom_memE(3)[OF A(2)] b c unfolding a_kernel_def' by auto | |
| 439 | hence "f (c \<oplus> b) = f a" | |
| 440 | using ring_hom_memE(3)[OF A(2) minus_closed[of a b], of b] a b by algebra | |
| 441 | moreover have "g (c \<oplus> b) = g b" | |
| 442 | using ring_hom_memE(1,3)[OF A(3)] b(1) c ring.ring_simprules(8)[OF A(1)] | |
| 443 | unfolding a_kernel_def' by auto | |
| 444 | ultimately have "(\<lambda>a. (f a, g a)) (c \<oplus> b) = t" and "c \<oplus> b \<in> carrier R" | |
| 445 | using a b c unfolding a_kernel_def' by auto | |
| 446 | thus "t \<in> (\<lambda>a. (f a, g a)) ` carrier R" | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 447 | by blast | 
| 81438 | 448 | qed | 
| 449 | qed | |
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 450 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 451 | let ?map_Quot = "\<lambda>I n. map (\<lambda>i. R Quot (I i)) [0..< Suc n]" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 452 | let ?DirProd = "\<lambda>I n. RDirProd_list (?map_Quot I n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 453 | let ?proj = "\<lambda>I n. canonical_proj_ext I n" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 454 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 455 | case (Suc n) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 456 | interpret I: ideal "I 0" R + Inter: ideal "\<Inter>i \<le> n. I (Suc i)" R | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 457 |     using i_Intersect[of "(\<lambda>i. I (Suc i)) ` {..n}"] Suc(2) by auto
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 458 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 459 | have map_simp: "?map_Quot I (Suc n) = (R Quot I 0) # (?map_Quot (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 460 | by (induct n) (auto) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 461 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 462 | have IH: "(?proj (\<lambda>i. I (Suc i)) n) ` carrier R = carrier (?DirProd (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 463 | and ring: "ring (?DirProd (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 464 | and hom: "?proj (\<lambda>i. I (Suc i)) n \<in> ring_hom R (?DirProd (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 465 | using RDirProd_Quot_list_is_ring[of n "\<lambda>i. I (Suc i)"] Suc(1)[of "\<lambda>i. I (Suc i)"] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 466 | canonical_proj_ext_is_hom[of n "\<lambda>i. I (Suc i)"] Suc(2-3) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 467 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 468 | have ker: "a_kernel R (?DirProd (\<lambda>i. I (Suc i)) n) (?proj (\<lambda>i. I (Suc i)) n) = (\<Inter>i \<le> n. I (Suc i))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 469 | using canonical_proj_ext_ker[of n "\<lambda>i. I (Suc i)"] Suc(2) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 470 | have carrier_Quot: "carrier (R Quot (I 0)) = (\<lambda>a. (I 0) +> a) ` carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 471 | by (auto simp add: RDirProd_list_carrier FactRing_def A_RCOSETS_def') | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 472 | have ring: "ring (?DirProd (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 473 | and hom: "?proj (\<lambda>i. I (Suc i)) n \<in> ring_hom R (?DirProd (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 474 | using RDirProd_Quot_list_is_ring[of n "\<lambda>i. I (Suc i)"] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 475 | canonical_proj_ext_is_hom[of n "\<lambda>i. I (Suc i)"] Suc(2) by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 476 | have "carrier (R Quot (I 0)) \<subseteq> (\<lambda>a. (I 0) +> a) ` (\<Inter>i \<le> n. I (Suc i))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 477 | proof | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 478 |     have "(\<Inter>i \<in> {Suc 0.. Suc n}. I i) <+> (I 0) = carrier R"
 | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 479 | using inter_plus_ideal_eq_carrier_arbitrary[of n I 0] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 480 | by (simp add: Suc(2-3) atLeast1_atMost_eq_remove0) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 481 | hence eq_carrier: "(I 0) <+> (\<Inter>i \<le> n. I (Suc i)) = carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 482 | using set_add_comm[OF I.a_subset Inter.a_subset] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 483 | by (metis INT_extend_simps(10) atMost_atLeast0 image_Suc_atLeastAtMost) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 484 | |
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 485 | fix b assume "b \<in> carrier (R Quot (I 0))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 486 | hence "(b, (\<Inter>i \<le> n. I (Suc i))) \<in> carrier (R Quot I 0) \<times> carrier (R Quot (\<Inter>i\<le>n. I (Suc i)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 487 | using ring.ring_simprules(2)[OF Inter.quotient_is_ring] by (simp add: FactRing_def) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 488 | then obtain s | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 489 | where "s \<in> carrier R" "(canonical_proj (I 0) (\<Inter>i \<le> n. I (Suc i))) s = (b, (\<Inter>i \<le> n. I (Suc i)))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 490 | using canonical_proj_is_surj[OF I.is_ideal Inter.is_ideal eq_carrier] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 491 | unfolding RDirProd_carrier by (metis (no_types, lifting) imageE) | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 492 | hence "s \<in> (\<Inter>i \<le> n. I (Suc i))" and "(\<lambda>a. (I 0) +> a) s = b" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 493 | using Inter.rcos_const_imp_mem by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 494 | thus "b \<in> (\<lambda>a. (I 0) +> a) ` (\<Inter>i \<le> n. I (Suc i))" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 495 | by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 496 | qed | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 497 | hence "(\<lambda>a. ((I 0) +> a, ?proj (\<lambda>i. I (Suc i)) n a)) ` carrier R = | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 498 | carrier (R Quot (I 0)) \<times> carrier (?DirProd (\<lambda>i. I (Suc i)) n)" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 499 | using aux_lemma[OF ring I.rcos_ring_hom hom] unfolding carrier_Quot ker IH by simp | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 500 | moreover show ?case | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 501 | unfolding map_simp RDirProd_list_carrier sym[OF calculation(1)] canonical_proj_ext_tl by auto | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 502 | qed | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 503 | |
| 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 504 | theorem (in cring) chinese_remainder: | 
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 505 | assumes "\<And>i. i \<le> n \<Longrightarrow> ideal (I i) R" and "\<And>i j. \<lbrakk> i \<le> n; j \<le> n \<rbrakk> \<Longrightarrow> i \<noteq> j \<Longrightarrow> I i <+> I j = carrier R" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 506 | shows "R Quot (\<Inter>i \<le> n. I i) \<simeq> RDirProd_list (map (\<lambda>i. R Quot (I i)) [0..< Suc n])" | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 507 | using ring_hom_ring.FactRing_iso[OF canonical_proj_ext_ring_hom, of n I] | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 508 | canonical_proj_ext_is_surj[of n I] canonical_proj_ext_ker[of n I] assms | 
| 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 509 | by auto | 
| 68569 
c64319959bab
Lots of new algebra theories by Martin Baillon and Paulo Emílio de Vilhena
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 510 | |
| 69122 
1b5178abaf97
updates to Algebra from Baillon and de Vilhena
 paulson <lp15@cam.ac.uk> parents: 
68975diff
changeset | 511 | end |