| author | wenzelm | 
| Thu, 08 Nov 2018 22:29:09 +0100 | |
| changeset 69272 | 15e9ed5b28fb | 
| parent 69064 | 5840724b1d71 | 
| child 69884 | dec7cc38a5dc | 
| permissions | -rw-r--r-- | 
| 65435 | 1 | (* Title: HOL/Computational_Algebra/Euclidean_Algorithm.thy | 
| 64784 
5cb5e7ecb284
reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
 haftmann parents: 
64592diff
changeset | 2 | Author: Manuel Eberl, TU Muenchen | 
| 
5cb5e7ecb284
reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
 haftmann parents: 
64592diff
changeset | 3 | *) | 
| 58023 | 4 | |
| 64784 
5cb5e7ecb284
reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
 haftmann parents: 
64592diff
changeset | 5 | section \<open>Abstract euclidean algorithm in euclidean (semi)rings\<close> | 
| 58023 | 6 | |
| 7 | theory Euclidean_Algorithm | |
| 65417 | 8 | imports Factorial_Ring | 
| 58023 | 9 | begin | 
| 10 | ||
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 11 | subsection \<open>Generic construction of the (simple) euclidean algorithm\<close> | 
| 66817 | 12 | |
| 13 | class normalization_euclidean_semiring = euclidean_semiring + normalization_semidom | |
| 64784 
5cb5e7ecb284
reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
 haftmann parents: 
64592diff
changeset | 14 | begin | 
| 63498 | 15 | |
| 66817 | 16 | lemma euclidean_size_normalize [simp]: | 
| 17 | "euclidean_size (normalize a) = euclidean_size a" | |
| 18 | proof (cases "a = 0") | |
| 19 | case True | |
| 20 | then show ?thesis | |
| 21 | by simp | |
| 22 | next | |
| 23 | case [simp]: False | |
| 24 | have "euclidean_size (normalize a) \<le> euclidean_size (normalize a * unit_factor a)" | |
| 25 | by (rule size_mult_mono) simp | |
| 26 | moreover have "euclidean_size a \<le> euclidean_size (a * (1 div unit_factor a))" | |
| 27 | by (rule size_mult_mono) simp | |
| 28 | ultimately show ?thesis | |
| 29 | by simp | |
| 30 | qed | |
| 31 | ||
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 32 | context | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 33 | begin | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 34 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 35 | qualified function gcd :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 36 | where "gcd a b = (if b = 0 then normalize a else gcd b (a mod b))" | 
| 60572 
718b1ba06429
streamlined definitions and primitive lemma of euclidean algorithm, including code generation
 haftmann parents: 
60571diff
changeset | 37 | by pat_completeness simp | 
| 60569 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 38 | termination | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 39 | by (relation "measure (euclidean_size \<circ> snd)") (simp_all add: mod_size_less) | 
| 58023 | 40 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 41 | declare gcd.simps [simp del] | 
| 58023 | 42 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 43 | lemma eucl_induct [case_names zero mod]: | 
| 60569 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 44 | assumes H1: "\<And>b. P b 0" | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 45 | and H2: "\<And>a b. b \<noteq> 0 \<Longrightarrow> P b (a mod b) \<Longrightarrow> P a b" | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 46 | shows "P a b" | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 47 | proof (induct a b rule: gcd.induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 48 | case (1 a b) | 
| 60569 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 49 | show ?case | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 50 | proof (cases "b = 0") | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 51 | case True then show "P a b" by simp (rule H1) | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 52 | next | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 53 | case False | 
| 60600 
87fbfea0bd0a
simplified termination criterion for euclidean algorithm (again)
 haftmann parents: 
60599diff
changeset | 54 | then have "P b (a mod b)" | 
| 
87fbfea0bd0a
simplified termination criterion for euclidean algorithm (again)
 haftmann parents: 
60599diff
changeset | 55 | by (rule "1.hyps") | 
| 60569 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 56 | with \<open>b \<noteq> 0\<close> show "P a b" | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 57 | by (blast intro: H2) | 
| 
f2f1f6860959
generalized to definition from literature, which covers also polynomials
 haftmann parents: 
60526diff
changeset | 58 | qed | 
| 58023 | 59 | qed | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 60 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 61 | qualified lemma gcd_0: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 62 | "gcd a 0 = normalize a" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 63 | by (simp add: gcd.simps [of a 0]) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 64 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 65 | qualified lemma gcd_mod: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 66 | "a \<noteq> 0 \<Longrightarrow> gcd a (b mod a) = gcd b a" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 67 | by (simp add: gcd.simps [of b 0] gcd.simps [of b a]) | 
| 58023 | 68 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 69 | qualified definition lcm :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 70 | where "lcm a b = normalize (a * b) div gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 71 | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 72 | qualified definition Lcm :: "'a set \<Rightarrow> 'a" \<comment> \<open>Somewhat complicated definition of Lcm that has the advantage of working | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 73 | for infinite sets as well\<close> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 74 | where | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 75 | [code del]: "Lcm A = (if \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) then | 
| 60430 
ce559c850a27
standardized algebraic conventions: prefer a, b, c over x, y, z
 haftmann parents: 
59061diff
changeset | 76 | let l = SOME l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = | 
| 
ce559c850a27
standardized algebraic conventions: prefer a, b, c over x, y, z
 haftmann parents: 
59061diff
changeset | 77 | (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n) | 
| 60634 | 78 | in normalize l | 
| 58023 | 79 | else 0)" | 
| 80 | ||
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 81 | qualified definition Gcd :: "'a set \<Rightarrow> 'a" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 82 |   where [code del]: "Gcd A = Lcm {d. \<forall>a\<in>A. d dvd a}"
 | 
| 62428 
4d5fbec92bb1
Fixed code equations for Gcd/Lcm
 Manuel Eberl <eberlm@in.tum.de> parents: 
62425diff
changeset | 83 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 84 | end | 
| 60572 
718b1ba06429
streamlined definitions and primitive lemma of euclidean algorithm, including code generation
 haftmann parents: 
60571diff
changeset | 85 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 86 | lemma semiring_gcd: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 87 | "class.semiring_gcd one zero times gcd lcm | 
| 64848 
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
 haftmann parents: 
64786diff
changeset | 88 | divide plus minus unit_factor normalize" | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 89 | proof | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 90 | show "gcd a b dvd a" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 91 | and "gcd a b dvd b" for a b | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 92 | by (induct a b rule: eucl_induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 93 | (simp_all add: local.gcd_0 local.gcd_mod dvd_mod_iff) | 
| 62422 | 94 | next | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 95 | show "c dvd a \<Longrightarrow> c dvd b \<Longrightarrow> c dvd gcd a b" for a b c | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 96 | proof (induct a b rule: eucl_induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 97 | case (zero a) from \<open>c dvd a\<close> show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 98 | by (rule dvd_trans) (simp add: local.gcd_0) | 
| 62422 | 99 | next | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 100 | case (mod a b) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 101 | then show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 102 | by (simp add: local.gcd_mod dvd_mod_iff) | 
| 62422 | 103 | qed | 
| 104 | next | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 105 | show "normalize (gcd a b) = gcd a b" for a b | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 106 | by (induct a b rule: eucl_induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 107 | (simp_all add: local.gcd_0 local.gcd_mod) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 108 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 109 | show "lcm a b = normalize (a * b) div gcd a b" for a b | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 110 | by (fact local.lcm_def) | 
| 63498 | 111 | qed | 
| 112 | ||
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 113 | interpretation semiring_gcd one zero times gcd lcm | 
| 64848 
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
 haftmann parents: 
64786diff
changeset | 114 | divide plus minus unit_factor normalize | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 115 | by (fact semiring_gcd) | 
| 63498 | 116 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 117 | lemma semiring_Gcd: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 118 | "class.semiring_Gcd one zero times gcd lcm Gcd Lcm | 
| 64848 
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
 haftmann parents: 
64786diff
changeset | 119 | divide plus minus unit_factor normalize" | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 120 | proof - | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 121 | show ?thesis | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 122 | proof | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 123 | have "(\<forall>a\<in>A. a dvd Lcm A) \<and> (\<forall>b. (\<forall>a\<in>A. a dvd b) \<longrightarrow> Lcm A dvd b)" for A | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 124 | proof (cases "\<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l)") | 
| 63498 | 125 | case False | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 126 | then have "Lcm A = 0" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 127 | by (auto simp add: local.Lcm_def) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 128 | with False show ?thesis | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 129 | by auto | 
| 63498 | 130 | next | 
| 131 | case True | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 132 | then obtain l\<^sub>0 where l\<^sub>0_props: "l\<^sub>0 \<noteq> 0" "\<forall>a\<in>A. a dvd l\<^sub>0" by blast | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 133 | define n where "n = (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 134 | define l where "l = (SOME l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 135 | have "\<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 136 | apply (subst n_def) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 137 | apply (rule LeastI [of _ "euclidean_size l\<^sub>0"]) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 138 | apply (rule exI [of _ l\<^sub>0]) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 139 | apply (simp add: l\<^sub>0_props) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 140 | done | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 141 | from someI_ex [OF this] have "l \<noteq> 0" and "\<forall>a\<in>A. a dvd l" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 142 | and "euclidean_size l = n" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 143 | unfolding l_def by simp_all | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 144 |       {
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 145 | fix l' assume "\<forall>a\<in>A. a dvd l'" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 146 | with \<open>\<forall>a\<in>A. a dvd l\<close> have "\<forall>a\<in>A. a dvd gcd l l'" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 147 | by (auto intro: gcd_greatest) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 148 | moreover from \<open>l \<noteq> 0\<close> have "gcd l l' \<noteq> 0" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 149 | by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 150 | ultimately have "\<exists>b. b \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd b) \<and> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 151 | euclidean_size b = euclidean_size (gcd l l')" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 152 | by (intro exI [of _ "gcd l l'"], auto) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 153 | then have "euclidean_size (gcd l l') \<ge> n" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 154 | by (subst n_def) (rule Least_le) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 155 | moreover have "euclidean_size (gcd l l') \<le> n" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 156 | proof - | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 157 | have "gcd l l' dvd l" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 158 | by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 159 | then obtain a where "l = gcd l l' * a" .. | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 160 | with \<open>l \<noteq> 0\<close> have "a \<noteq> 0" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 161 | by auto | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 162 | hence "euclidean_size (gcd l l') \<le> euclidean_size (gcd l l' * a)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 163 | by (rule size_mult_mono) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 164 | also have "gcd l l' * a = l" using \<open>l = gcd l l' * a\<close> .. | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 165 | also note \<open>euclidean_size l = n\<close> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 166 | finally show "euclidean_size (gcd l l') \<le> n" . | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 167 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 168 | ultimately have *: "euclidean_size l = euclidean_size (gcd l l')" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 169 | by (intro le_antisym, simp_all add: \<open>euclidean_size l = n\<close>) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 170 | from \<open>l \<noteq> 0\<close> have "l dvd gcd l l'" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 171 | by (rule dvd_euclidean_size_eq_imp_dvd) (auto simp add: *) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 172 | hence "l dvd l'" by (rule dvd_trans [OF _ gcd_dvd2]) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 173 | } | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 174 | with \<open>\<forall>a\<in>A. a dvd l\<close> and \<open>l \<noteq> 0\<close> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 175 | have "(\<forall>a\<in>A. a dvd normalize l) \<and> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 176 | (\<forall>l'. (\<forall>a\<in>A. a dvd l') \<longrightarrow> normalize l dvd l')" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 177 | by auto | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 178 | also from True have "normalize l = Lcm A" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 179 | by (simp add: local.Lcm_def Let_def n_def l_def) | 
| 63498 | 180 | finally show ?thesis . | 
| 181 | qed | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 182 | then show dvd_Lcm: "a \<in> A \<Longrightarrow> a dvd Lcm A" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 183 | and Lcm_least: "(\<And>a. a \<in> A \<Longrightarrow> a dvd b) \<Longrightarrow> Lcm A dvd b" for A and a b | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 184 | by auto | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 185 | show "a \<in> A \<Longrightarrow> Gcd A dvd a" for A and a | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 186 | by (auto simp add: local.Gcd_def intro: Lcm_least) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 187 | show "(\<And>a. a \<in> A \<Longrightarrow> b dvd a) \<Longrightarrow> b dvd Gcd A" for A and b | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 188 | by (auto simp add: local.Gcd_def intro: dvd_Lcm) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 189 | show [simp]: "normalize (Lcm A) = Lcm A" for A | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 190 | by (simp add: local.Lcm_def) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 191 | show "normalize (Gcd A) = Gcd A" for A | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 192 | by (simp add: local.Gcd_def) | 
| 62442 
26e4be6a680f
More efficient Extended Euclidean Algorithm
 Manuel Eberl <eberlm@in.tum.de> parents: 
62429diff
changeset | 193 | qed | 
| 
26e4be6a680f
More efficient Extended Euclidean Algorithm
 Manuel Eberl <eberlm@in.tum.de> parents: 
62429diff
changeset | 194 | qed | 
| 
26e4be6a680f
More efficient Extended Euclidean Algorithm
 Manuel Eberl <eberlm@in.tum.de> parents: 
62429diff
changeset | 195 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 196 | interpretation semiring_Gcd one zero times gcd lcm Gcd Lcm | 
| 64848 
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
 haftmann parents: 
64786diff
changeset | 197 | divide plus minus unit_factor normalize | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 198 | by (fact semiring_Gcd) | 
| 60598 
78ca5674c66a
rings follow immediately their corresponding semirings
 haftmann parents: 
60582diff
changeset | 199 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 200 | subclass factorial_semiring | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 201 | proof - | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 202 | show "class.factorial_semiring divide plus minus zero times one | 
| 64848 
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
 haftmann parents: 
64786diff
changeset | 203 | unit_factor normalize" | 
| 64911 | 204 | proof (standard, rule factorial_semiring_altI_aux) \<comment> \<open>FIXME rule\<close> | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 205 | fix x assume "x \<noteq> 0" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 206 |     thus "finite {p. p dvd x \<and> normalize p = p}"
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 207 | proof (induction "euclidean_size x" arbitrary: x rule: less_induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 208 | case (less x) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 209 | show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 210 | proof (cases "\<exists>y. y dvd x \<and> \<not>x dvd y \<and> \<not>is_unit y") | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 211 | case False | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 212 |         have "{p. p dvd x \<and> normalize p = p} \<subseteq> {1, normalize x}"
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 213 | proof | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 214 |           fix p assume p: "p \<in> {p. p dvd x \<and> normalize p = p}"
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 215 | with False have "is_unit p \<or> x dvd p" by blast | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 216 |           thus "p \<in> {1, normalize x}"
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 217 | proof (elim disjE) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 218 | assume "is_unit p" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 219 | hence "normalize p = 1" by (simp add: is_unit_normalize) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 220 | with p show ?thesis by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 221 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 222 | assume "x dvd p" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 223 | with p have "normalize p = normalize x" by (intro associatedI) simp_all | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 224 | with p show ?thesis by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 225 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 226 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 227 | moreover have "finite \<dots>" by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 228 | ultimately show ?thesis by (rule finite_subset) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 229 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 230 | case True | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 231 | then obtain y where y: "y dvd x" "\<not>x dvd y" "\<not>is_unit y" by blast | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 232 | define z where "z = x div y" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 233 |         let ?fctrs = "\<lambda>x. {p. p dvd x \<and> normalize p = p}"
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 234 | from y have x: "x = y * z" by (simp add: z_def) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 235 | with less.prems have "y \<noteq> 0" "z \<noteq> 0" by auto | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 236 | have normalized_factors_product: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 237 |           "{p. p dvd a * b \<and> normalize p = p} = 
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 238 |              (\<lambda>(x,y). x * y) ` ({p. p dvd a \<and> normalize p = p} \<times> {p. p dvd b \<and> normalize p = p})" for a b
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 239 | proof safe | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 240 | fix p assume p: "p dvd a * b" "normalize p = p" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 241 | from dvd_productE[OF p(1)] guess x y . note xy = this | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 242 | define x' y' where "x' = normalize x" and "y' = normalize y" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 243 | have "p = x' * y'" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 244 | by (subst p(2) [symmetric]) (simp add: xy x'_def y'_def normalize_mult) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 245 | moreover from xy have "normalize x' = x'" "normalize y' = y'" "x' dvd a" "y' dvd b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 246 | by (simp_all add: x'_def y'_def) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 247 | ultimately show "p \<in> (\<lambda>(x, y). x * y) ` | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 248 |             ({p. p dvd a \<and> normalize p = p} \<times> {p. p dvd b \<and> normalize p = p})"
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 249 | by blast | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 250 | qed (auto simp: normalize_mult mult_dvd_mono) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 251 | from x y have "\<not>is_unit z" by (auto simp: mult_unit_dvd_iff) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 252 | have "?fctrs x = (\<lambda>(p,p'). p * p') ` (?fctrs y \<times> ?fctrs z)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 253 | by (subst x) (rule normalized_factors_product) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 254 | also have "\<not>y * z dvd y * 1" "\<not>y * z dvd 1 * z" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 255 | by (subst dvd_times_left_cancel_iff dvd_times_right_cancel_iff; fact)+ | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 256 | hence "finite ((\<lambda>(p,p'). p * p') ` (?fctrs y \<times> ?fctrs z))" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 257 | by (intro finite_imageI finite_cartesian_product less dvd_proper_imp_size_less) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 258 | (auto simp: x) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 259 | finally show ?thesis . | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 260 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 261 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 262 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 263 | fix p | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 264 | assume "irreducible p" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 265 | then show "prime_elem p" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 266 | by (rule irreducible_imp_prime_elem_gcd) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 267 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 268 | qed | 
| 60598 
78ca5674c66a
rings follow immediately their corresponding semirings
 haftmann parents: 
60582diff
changeset | 269 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 270 | lemma Gcd_eucl_set [code]: | 
| 64850 | 271 | "Gcd (set xs) = fold gcd xs 0" | 
| 272 | by (fact Gcd_set_eq_fold) | |
| 60598 
78ca5674c66a
rings follow immediately their corresponding semirings
 haftmann parents: 
60582diff
changeset | 273 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 274 | lemma Lcm_eucl_set [code]: | 
| 64850 | 275 | "Lcm (set xs) = fold lcm xs 1" | 
| 276 | by (fact Lcm_set_eq_fold) | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 277 | |
| 60598 
78ca5674c66a
rings follow immediately their corresponding semirings
 haftmann parents: 
60582diff
changeset | 278 | end | 
| 
78ca5674c66a
rings follow immediately their corresponding semirings
 haftmann parents: 
60582diff
changeset | 279 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 280 | hide_const (open) gcd lcm Gcd Lcm | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 281 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 282 | lemma prime_elem_int_abs_iff [simp]: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 283 | fixes p :: int | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 284 | shows "prime_elem \<bar>p\<bar> \<longleftrightarrow> prime_elem p" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 285 | using prime_elem_normalize_iff [of p] by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 286 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 287 | lemma prime_elem_int_minus_iff [simp]: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 288 | fixes p :: int | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 289 | shows "prime_elem (- p) \<longleftrightarrow> prime_elem p" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 290 | using prime_elem_normalize_iff [of "- p"] by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 291 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 292 | lemma prime_int_iff: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 293 | fixes p :: int | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 294 | shows "prime p \<longleftrightarrow> p > 0 \<and> prime_elem p" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 295 | by (auto simp add: prime_def dest: prime_elem_not_zeroI) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 296 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 297 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 298 | subsection \<open>The (simple) euclidean algorithm as gcd computation\<close> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 299 | |
| 66817 | 300 | class euclidean_semiring_gcd = normalization_euclidean_semiring + gcd + Gcd + | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 301 | assumes gcd_eucl: "Euclidean_Algorithm.gcd = GCD.gcd" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 302 | and lcm_eucl: "Euclidean_Algorithm.lcm = GCD.lcm" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 303 | assumes Gcd_eucl: "Euclidean_Algorithm.Gcd = GCD.Gcd" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 304 | and Lcm_eucl: "Euclidean_Algorithm.Lcm = GCD.Lcm" | 
| 58023 | 305 | begin | 
| 306 | ||
| 62422 | 307 | subclass semiring_gcd | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 308 | unfolding gcd_eucl [symmetric] lcm_eucl [symmetric] | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 309 | by (fact semiring_gcd) | 
| 58023 | 310 | |
| 62422 | 311 | subclass semiring_Gcd | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 312 | unfolding gcd_eucl [symmetric] lcm_eucl [symmetric] | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 313 | Gcd_eucl [symmetric] Lcm_eucl [symmetric] | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 314 | by (fact semiring_Gcd) | 
| 63498 | 315 | |
| 316 | subclass factorial_semiring_gcd | |
| 317 | proof | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 318 | show "gcd a b = gcd_factorial a b" for a b | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 319 | apply (rule sym) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 320 | apply (rule gcdI) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 321 | apply (fact gcd_lcm_factorial)+ | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 322 | done | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 323 | then show "lcm a b = lcm_factorial a b" for a b | 
| 63498 | 324 | by (simp add: lcm_factorial_gcd_factorial lcm_gcd) | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 325 | show "Gcd A = Gcd_factorial A" for A | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 326 | apply (rule sym) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 327 | apply (rule GcdI) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 328 | apply (fact gcd_lcm_factorial)+ | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 329 | done | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 330 | show "Lcm A = Lcm_factorial A" for A | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 331 | apply (rule sym) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 332 | apply (rule LcmI) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 333 | apply (fact gcd_lcm_factorial)+ | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 334 | done | 
| 63498 | 335 | qed | 
| 336 | ||
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 337 | lemma gcd_mod_right [simp]: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 338 | "a \<noteq> 0 \<Longrightarrow> gcd a (b mod a) = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 339 | unfolding gcd.commute [of a b] | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 340 | by (simp add: gcd_eucl [symmetric] local.gcd_mod) | 
| 58023 | 341 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 342 | lemma gcd_mod_left [simp]: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 343 | "b \<noteq> 0 \<Longrightarrow> gcd (a mod b) b = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 344 | by (drule gcd_mod_right [of _ a]) (simp add: gcd.commute) | 
| 58023 | 345 | |
| 346 | lemma euclidean_size_gcd_le1 [simp]: | |
| 347 | assumes "a \<noteq> 0" | |
| 348 | shows "euclidean_size (gcd a b) \<le> euclidean_size a" | |
| 349 | proof - | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 350 | from gcd_dvd1 obtain c where A: "a = gcd a b * c" .. | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 351 | with assms have "c \<noteq> 0" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 352 | by auto | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 353 | moreover from this | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 354 | have "euclidean_size (gcd a b) \<le> euclidean_size (gcd a b * c)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 355 | by (rule size_mult_mono) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 356 | with A show ?thesis | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 357 | by simp | 
| 58023 | 358 | qed | 
| 359 | ||
| 360 | lemma euclidean_size_gcd_le2 [simp]: | |
| 361 | "b \<noteq> 0 \<Longrightarrow> euclidean_size (gcd a b) \<le> euclidean_size b" | |
| 362 | by (subst gcd.commute, rule euclidean_size_gcd_le1) | |
| 363 | ||
| 364 | lemma euclidean_size_gcd_less1: | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 365 | assumes "a \<noteq> 0" and "\<not> a dvd b" | 
| 58023 | 366 | shows "euclidean_size (gcd a b) < euclidean_size a" | 
| 367 | proof (rule ccontr) | |
| 368 | assume "\<not>euclidean_size (gcd a b) < euclidean_size a" | |
| 62422 | 369 | with \<open>a \<noteq> 0\<close> have A: "euclidean_size (gcd a b) = euclidean_size a" | 
| 58023 | 370 | by (intro le_antisym, simp_all) | 
| 62422 | 371 | have "a dvd gcd a b" | 
| 372 | by (rule dvd_euclidean_size_eq_imp_dvd) (simp_all add: assms A) | |
| 373 | hence "a dvd b" using dvd_gcdD2 by blast | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 374 | with \<open>\<not> a dvd b\<close> show False by contradiction | 
| 58023 | 375 | qed | 
| 376 | ||
| 377 | lemma euclidean_size_gcd_less2: | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 378 | assumes "b \<noteq> 0" and "\<not> b dvd a" | 
| 58023 | 379 | shows "euclidean_size (gcd a b) < euclidean_size b" | 
| 380 | using assms by (subst gcd.commute, rule euclidean_size_gcd_less1) | |
| 381 | ||
| 382 | lemma euclidean_size_lcm_le1: | |
| 383 | assumes "a \<noteq> 0" and "b \<noteq> 0" | |
| 384 | shows "euclidean_size a \<le> euclidean_size (lcm a b)" | |
| 385 | proof - | |
| 60690 | 386 | have "a dvd lcm a b" by (rule dvd_lcm1) | 
| 387 | then obtain c where A: "lcm a b = a * c" .. | |
| 62429 
25271ff79171
Tuned Euclidean Rings/GCD rings
 Manuel Eberl <eberlm@in.tum.de> parents: 
62428diff
changeset | 388 | with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> have "c \<noteq> 0" by (auto simp: lcm_eq_0_iff) | 
| 58023 | 389 | then show ?thesis by (subst A, intro size_mult_mono) | 
| 390 | qed | |
| 391 | ||
| 392 | lemma euclidean_size_lcm_le2: | |
| 393 | "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> euclidean_size b \<le> euclidean_size (lcm a b)" | |
| 394 | using euclidean_size_lcm_le1 [of b a] by (simp add: ac_simps) | |
| 395 | ||
| 396 | lemma euclidean_size_lcm_less1: | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 397 | assumes "b \<noteq> 0" and "\<not> b dvd a" | 
| 58023 | 398 | shows "euclidean_size a < euclidean_size (lcm a b)" | 
| 399 | proof (rule ccontr) | |
| 400 | from assms have "a \<noteq> 0" by auto | |
| 401 | assume "\<not>euclidean_size a < euclidean_size (lcm a b)" | |
| 60526 | 402 | with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> have "euclidean_size (lcm a b) = euclidean_size a" | 
| 58023 | 403 | by (intro le_antisym, simp, intro euclidean_size_lcm_le1) | 
| 404 | with assms have "lcm a b dvd a" | |
| 62429 
25271ff79171
Tuned Euclidean Rings/GCD rings
 Manuel Eberl <eberlm@in.tum.de> parents: 
62428diff
changeset | 405 | by (rule_tac dvd_euclidean_size_eq_imp_dvd) (auto simp: lcm_eq_0_iff) | 
| 62422 | 406 | hence "b dvd a" by (rule lcm_dvdD2) | 
| 60526 | 407 | with \<open>\<not>b dvd a\<close> show False by contradiction | 
| 58023 | 408 | qed | 
| 409 | ||
| 410 | lemma euclidean_size_lcm_less2: | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 411 | assumes "a \<noteq> 0" and "\<not> a dvd b" | 
| 58023 | 412 | shows "euclidean_size b < euclidean_size (lcm a b)" | 
| 413 | using assms euclidean_size_lcm_less1 [of a b] by (simp add: ac_simps) | |
| 414 | ||
| 415 | end | |
| 416 | ||
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 417 | lemma factorial_euclidean_semiring_gcdI: | 
| 66817 | 418 |   "OFCLASS('a::{factorial_semiring_gcd, normalization_euclidean_semiring}, euclidean_semiring_gcd_class)"
 | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 419 | proof | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 420 | interpret semiring_Gcd 1 0 times | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 421 | Euclidean_Algorithm.gcd Euclidean_Algorithm.lcm | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 422 | Euclidean_Algorithm.Gcd Euclidean_Algorithm.Lcm | 
| 64848 
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
 haftmann parents: 
64786diff
changeset | 423 | divide plus minus unit_factor normalize | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67443diff
changeset | 424 | rewrites "dvd.dvd (*) = Rings.dvd" | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 425 | by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 426 | show [simp]: "Euclidean_Algorithm.gcd = (gcd :: 'a \<Rightarrow> _)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 427 | proof (rule ext)+ | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 428 | fix a b :: 'a | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 429 | show "Euclidean_Algorithm.gcd a b = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 430 | proof (induct a b rule: eucl_induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 431 | case zero | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 432 | then show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 433 | by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 434 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 435 | case (mod a b) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 436 | moreover have "gcd b (a mod b) = gcd b a" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 437 | using GCD.gcd_add_mult [of b "a div b" "a mod b", symmetric] | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 438 | by (simp add: div_mult_mod_eq) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 439 | ultimately show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 440 | by (simp add: Euclidean_Algorithm.gcd_mod ac_simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 441 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 442 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 443 | show [simp]: "Euclidean_Algorithm.Lcm = (Lcm :: 'a set \<Rightarrow> _)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 444 | by (auto intro!: Lcm_eqI GCD.dvd_Lcm GCD.Lcm_least) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 445 | show "Euclidean_Algorithm.lcm = (lcm :: 'a \<Rightarrow> _)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 446 | by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 447 | show "Euclidean_Algorithm.Gcd = (Gcd :: 'a set \<Rightarrow> _)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 448 | by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 449 | qed | 
| 63498 | 450 | |
| 58023 | 451 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 452 | subsection \<open>The extended euclidean algorithm\<close> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 453 | |
| 58023 | 454 | class euclidean_ring_gcd = euclidean_semiring_gcd + idom | 
| 455 | begin | |
| 456 | ||
| 457 | subclass euclidean_ring .. | |
| 60439 
b765e08f8bc0
proper subclass instances for existing gcd (semi)rings
 haftmann parents: 
60438diff
changeset | 458 | subclass ring_gcd .. | 
| 63498 | 459 | subclass factorial_ring_gcd .. | 
| 60439 
b765e08f8bc0
proper subclass instances for existing gcd (semi)rings
 haftmann parents: 
60438diff
changeset | 460 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 461 | function euclid_ext_aux :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'a) \<times> 'a"
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 462 | where "euclid_ext_aux s' s t' t r' r = ( | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 463 | if r = 0 then let c = 1 div unit_factor r' in ((s' * c, t' * c), normalize r') | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 464 | else let q = r' div r | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 465 | in euclid_ext_aux s (s' - q * s) t (t' - q * t) r (r' mod r))" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 466 | by auto | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 467 | termination | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 468 | by (relation "measure (\<lambda>(_, _, _, _, _, b). euclidean_size b)") | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 469 | (simp_all add: mod_size_less) | 
| 62442 
26e4be6a680f
More efficient Extended Euclidean Algorithm
 Manuel Eberl <eberlm@in.tum.de> parents: 
62429diff
changeset | 470 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 471 | abbreviation (input) euclid_ext :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'a) \<times> 'a"
 | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 472 | where "euclid_ext \<equiv> euclid_ext_aux 1 0 0 1" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 473 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 474 | lemma | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 475 | assumes "gcd r' r = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 476 | assumes "s' * a + t' * b = r'" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 477 | assumes "s * a + t * b = r" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 478 | assumes "euclid_ext_aux s' s t' t r' r = ((x, y), c)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 479 | shows euclid_ext_aux_eq_gcd: "c = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 480 | and euclid_ext_aux_bezout: "x * a + y * b = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 481 | proof - | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 482 | have "case euclid_ext_aux s' s t' t r' r of ((x, y), c) \<Rightarrow> | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 483 | x * a + y * b = c \<and> c = gcd a b" (is "?P (euclid_ext_aux s' s t' t r' r)") | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 484 | using assms(1-3) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 485 | proof (induction s' s t' t r' r rule: euclid_ext_aux.induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 486 | case (1 s' s t' t r' r) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 487 | show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 488 | proof (cases "r = 0") | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 489 | case True | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 490 | hence "euclid_ext_aux s' s t' t r' r = | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 491 | ((s' div unit_factor r', t' div unit_factor r'), normalize r')" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 492 | by (subst euclid_ext_aux.simps) (simp add: Let_def) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 493 | also have "?P \<dots>" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 494 | proof safe | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 495 | have "s' div unit_factor r' * a + t' div unit_factor r' * b = | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 496 | (s' * a + t' * b) div unit_factor r'" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 497 | by (cases "r' = 0") (simp_all add: unit_div_commute) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 498 | also have "s' * a + t' * b = r'" by fact | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 499 | also have "\<dots> div unit_factor r' = normalize r'" by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 500 | finally show "s' div unit_factor r' * a + t' div unit_factor r' * b = normalize r'" . | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 501 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 502 | from "1.prems" True show "normalize r' = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 503 | by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 504 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 505 | finally show ?thesis . | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 506 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 507 | case False | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 508 | hence "euclid_ext_aux s' s t' t r' r = | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 509 | euclid_ext_aux s (s' - r' div r * s) t (t' - r' div r * t) r (r' mod r)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 510 | by (subst euclid_ext_aux.simps) (simp add: Let_def) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 511 | also from "1.prems" False have "?P \<dots>" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 512 | proof (intro "1.IH") | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 513 | have "(s' - r' div r * s) * a + (t' - r' div r * t) * b = | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 514 | (s' * a + t' * b) - r' div r * (s * a + t * b)" by (simp add: algebra_simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 515 | also have "s' * a + t' * b = r'" by fact | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 516 | also have "s * a + t * b = r" by fact | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 517 | also have "r' - r' div r * r = r' mod r" using div_mult_mod_eq [of r' r] | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 518 | by (simp add: algebra_simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 519 | finally show "(s' - r' div r * s) * a + (t' - r' div r * t) * b = r' mod r" . | 
| 66817 | 520 | qed (auto simp: algebra_simps minus_mod_eq_div_mult [symmetric] gcd.commute) | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 521 | finally show ?thesis . | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 522 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 523 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 524 | with assms(4) show "c = gcd a b" "x * a + y * b = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 525 | by simp_all | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 526 | qed | 
| 60572 
718b1ba06429
streamlined definitions and primitive lemma of euclidean algorithm, including code generation
 haftmann parents: 
60571diff
changeset | 527 | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 528 | declare euclid_ext_aux.simps [simp del] | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 529 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 530 | definition bezout_coefficients :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<times> 'a" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 531 | where [code]: "bezout_coefficients a b = fst (euclid_ext a b)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 532 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 533 | lemma bezout_coefficients_0: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 534 | "bezout_coefficients a 0 = (1 div unit_factor a, 0)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 535 | by (simp add: bezout_coefficients_def euclid_ext_aux.simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 536 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 537 | lemma bezout_coefficients_left_0: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 538 | "bezout_coefficients 0 a = (0, 1 div unit_factor a)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 539 | by (simp add: bezout_coefficients_def euclid_ext_aux.simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 540 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 541 | lemma bezout_coefficients: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 542 | assumes "bezout_coefficients a b = (x, y)" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 543 | shows "x * a + y * b = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 544 | using assms by (simp add: bezout_coefficients_def | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 545 | euclid_ext_aux_bezout [of a b a b 1 0 0 1 x y] prod_eq_iff) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 546 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 547 | lemma bezout_coefficients_fst_snd: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 548 | "fst (bezout_coefficients a b) * a + snd (bezout_coefficients a b) * b = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 549 | by (rule bezout_coefficients) simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 550 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 551 | lemma euclid_ext_eq [simp]: | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 552 | "euclid_ext a b = (bezout_coefficients a b, gcd a b)" (is "?p = ?q") | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 553 | proof | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 554 | show "fst ?p = fst ?q" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 555 | by (simp add: bezout_coefficients_def) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 556 | have "snd (euclid_ext_aux 1 0 0 1 a b) = gcd a b" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 557 | by (rule euclid_ext_aux_eq_gcd [of a b a b 1 0 0 1]) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 558 | (simp_all add: prod_eq_iff) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 559 | then show "snd ?p = snd ?q" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 560 | by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 561 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 562 | |
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 563 | declare euclid_ext_eq [symmetric, code_unfold] | 
| 60572 
718b1ba06429
streamlined definitions and primitive lemma of euclidean algorithm, including code generation
 haftmann parents: 
60571diff
changeset | 564 | |
| 
718b1ba06429
streamlined definitions and primitive lemma of euclidean algorithm, including code generation
 haftmann parents: 
60571diff
changeset | 565 | end | 
| 58023 | 566 | |
| 567 | ||
| 60572 
718b1ba06429
streamlined definitions and primitive lemma of euclidean algorithm, including code generation
 haftmann parents: 
60571diff
changeset | 568 | subsection \<open>Typical instances\<close> | 
| 58023 | 569 | |
| 66817 | 570 | instance nat :: normalization_euclidean_semiring .. | 
| 571 | ||
| 62422 | 572 | instance nat :: euclidean_semiring_gcd | 
| 573 | proof | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 574 | interpret semiring_Gcd 1 0 times | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 575 | "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 576 | "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm" | 
| 64848 
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
 haftmann parents: 
64786diff
changeset | 577 | divide plus minus unit_factor normalize | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67443diff
changeset | 578 | rewrites "dvd.dvd (*) = Rings.dvd" | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 579 | by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 580 | show [simp]: "(Euclidean_Algorithm.gcd :: nat \<Rightarrow> _) = gcd" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 581 | proof (rule ext)+ | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 582 | fix m n :: nat | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 583 | show "Euclidean_Algorithm.gcd m n = gcd m n" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 584 | proof (induct m n rule: eucl_induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 585 | case zero | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 586 | then show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 587 | by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 588 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 589 | case (mod m n) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 590 | then have "gcd n (m mod n) = gcd n m" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 591 | using gcd_nat.simps [of m n] by (simp add: ac_simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 592 | with mod show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 593 | by (simp add: Euclidean_Algorithm.gcd_mod ac_simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 594 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 595 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 596 | show [simp]: "(Euclidean_Algorithm.Lcm :: nat set \<Rightarrow> _) = Lcm" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 597 | by (auto intro!: ext Lcm_eqI) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 598 | show "(Euclidean_Algorithm.lcm :: nat \<Rightarrow> _) = lcm" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 599 | by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 600 | show "(Euclidean_Algorithm.Gcd :: nat set \<Rightarrow> _) = Gcd" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 601 | by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm) | 
| 62422 | 602 | qed | 
| 603 | ||
| 67167 
88d1c9d86f48
Moved analysis material from AFP
 Manuel Eberl <eberlm@in.tum.de> parents: 
66817diff
changeset | 604 | lemma prime_factorization_Suc_0 [simp]: "prime_factorization (Suc 0) = {#}"
 | 
| 
88d1c9d86f48
Moved analysis material from AFP
 Manuel Eberl <eberlm@in.tum.de> parents: 
66817diff
changeset | 605 | unfolding One_nat_def [symmetric] using prime_factorization_1 . | 
| 
88d1c9d86f48
Moved analysis material from AFP
 Manuel Eberl <eberlm@in.tum.de> parents: 
66817diff
changeset | 606 | |
| 66817 | 607 | instance int :: normalization_euclidean_semiring .. | 
| 608 | ||
| 62422 | 609 | instance int :: euclidean_ring_gcd | 
| 610 | proof | |
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 611 | interpret semiring_Gcd 1 0 times | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 612 | "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 613 | "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm" | 
| 64848 
c50db2128048
slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
 haftmann parents: 
64786diff
changeset | 614 | divide plus minus unit_factor normalize | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67443diff
changeset | 615 | rewrites "dvd.dvd (*) = Rings.dvd" | 
| 64786 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 616 | by (fact semiring_Gcd) (simp add: dvd.dvd_def dvd_def fun_eq_iff) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 617 | show [simp]: "(Euclidean_Algorithm.gcd :: int \<Rightarrow> _) = gcd" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 618 | proof (rule ext)+ | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 619 | fix k l :: int | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 620 | show "Euclidean_Algorithm.gcd k l = gcd k l" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 621 | proof (induct k l rule: eucl_induct) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 622 | case zero | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 623 | then show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 624 | by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 625 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 626 | case (mod k l) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 627 | have "gcd l (k mod l) = gcd l k" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 628 | proof (cases l "0::int" rule: linorder_cases) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 629 | case less | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 630 | then show ?thesis | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 631 | using gcd_non_0_int [of "- l" "- k"] by (simp add: ac_simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 632 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 633 | case equal | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 634 | with mod show ?thesis | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 635 | by simp | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 636 | next | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 637 | case greater | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 638 | then show ?thesis | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 639 | using gcd_non_0_int [of l k] by (simp add: ac_simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 640 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 641 | with mod show ?case | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 642 | by (simp add: Euclidean_Algorithm.gcd_mod ac_simps) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 643 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 644 | qed | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 645 | show [simp]: "(Euclidean_Algorithm.Lcm :: int set \<Rightarrow> _) = Lcm" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 646 | by (auto intro!: ext Lcm_eqI) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 647 | show "(Euclidean_Algorithm.lcm :: int \<Rightarrow> _) = lcm" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 648 | by (simp add: fun_eq_iff Euclidean_Algorithm.lcm_def semiring_gcd_class.lcm_gcd) | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 649 | show "(Euclidean_Algorithm.Gcd :: int set \<Rightarrow> _) = Gcd" | 
| 
340db65fd2c1
reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
 haftmann parents: 
64785diff
changeset | 650 | by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm) | 
| 62422 | 651 | qed | 
| 652 | ||
| 63924 | 653 | end |