src/HOL/Computational_Algebra/Euclidean_Algorithm.thy
author wenzelm
Mon, 09 Sep 2024 11:12:13 +0200
changeset 80821 eb383d50564b
parent 80084 173548e4d5d0
child 83001 157aaea4c42c
permissions -rw-r--r--
clarified signature: more explicit type "ops";
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
65435
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     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: 64592
diff changeset
     2
    Author:     Manuel Eberl, TU Muenchen
5cb5e7ecb284 reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
haftmann
parents: 64592
diff changeset
     3
*)
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
     4
64784
5cb5e7ecb284 reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
haftmann
parents: 64592
diff changeset
     5
section \<open>Abstract euclidean algorithm in euclidean (semi)rings\<close>
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
     6
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
     7
theory Euclidean_Algorithm
65417
fc41a5650fb1 session containing computational algebra
haftmann
parents: 65398
diff changeset
     8
  imports Factorial_Ring
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
     9
begin
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
    10
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    11
subsection \<open>Generic construction of the (simple) euclidean algorithm\<close>
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    12
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    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: 64592
diff changeset
    14
begin
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
    15
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    16
lemma euclidean_size_normalize [simp]:
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    17
  "euclidean_size (normalize a) = euclidean_size a"
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    18
proof (cases "a = 0")
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    19
  case True
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    20
  then show ?thesis
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    21
    by simp
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    22
next
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    23
  case [simp]: False
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    24
  have "euclidean_size (normalize a) \<le> euclidean_size (normalize a * unit_factor a)"
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    25
    by (rule size_mult_mono) simp
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    26
  moreover have "euclidean_size a \<le> euclidean_size (a * (1 div unit_factor a))"
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    27
    by (rule size_mult_mono) simp
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    28
  ultimately show ?thesis
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    29
    by simp
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    30
qed
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
    31
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    32
context
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    33
begin
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    34
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 60571
diff changeset
    37
  by pat_completeness simp
60569
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    38
termination
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    39
  by (relation "measure (euclidean_size \<circ> snd)") (simp_all add: mod_size_less)
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
    40
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    41
declare gcd.simps [simp del]
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
    42
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    43
lemma eucl_induct [case_names zero mod]:
60569
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    44
  assumes H1: "\<And>b. P b 0"
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff 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: 60526
diff changeset
    46
  shows "P a b"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
    48
  case (1 a b)
60569
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    49
  show ?case
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    50
  proof (cases "b = 0")
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff 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: 60526
diff changeset
    52
  next
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    53
    case False
60600
87fbfea0bd0a simplified termination criterion for euclidean algorithm (again)
haftmann
parents: 60599
diff changeset
    54
    then have "P b (a mod b)"
87fbfea0bd0a simplified termination criterion for euclidean algorithm (again)
haftmann
parents: 60599
diff changeset
    55
      by (rule "1.hyps")
60569
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    56
    with \<open>b \<noteq> 0\<close> show "P a b"
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    57
      by (blast intro: H2)
f2f1f6860959 generalized to definition from literature, which covers also polynomials
haftmann
parents: 60526
diff changeset
    58
  qed
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
    59
qed
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    60
  
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    61
qualified lemma gcd_0:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    62
  "gcd a 0 = normalize a"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
    64
  
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    65
qualified lemma gcd_mod:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
    67
  by (simp add: gcd.simps [of b 0] gcd.simps [of b a])
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
    68
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    69
qualified definition lcm :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
    70
  where "lcm a b = normalize (a * b div gcd a b)"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    71
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67399
diff 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: 64785
diff 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: 64785
diff changeset
    74
  where
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 59061
diff 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: 59061
diff changeset
    77
       (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>a\<in>A. a dvd l) \<and> euclidean_size l = n)
60634
e3b6e516608b separate (semi)ring with normalization
haftmann
parents: 60600
diff changeset
    78
       in normalize l 
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
    79
      else 0)"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
    80
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 62425
diff changeset
    83
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    84
end    
60572
718b1ba06429 streamlined definitions and primitive lemma of euclidean algorithm, including code generation
haftmann
parents: 60571
diff changeset
    85
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
    86
lemma semiring_gcd:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64786
diff 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: 64785
diff changeset
    89
proof
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
    93
      (simp_all add: local.gcd_0 local.gcd_mod dvd_mod_iff)
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
    94
next
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
    98
      by (rule dvd_trans) (simp add: local.gcd_0)
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
    99
  next
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   100
    case (mod a b)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   101
    then show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   102
      by (simp add: local.gcd_mod dvd_mod_iff)
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   103
  qed
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   104
next
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   108
next
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   109
  show "lcm a b = normalize (a * b div gcd a b)" for a b
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   110
    by (fact local.lcm_def)
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   111
qed
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   112
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64786
diff 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: 64785
diff changeset
   115
  by (fact semiring_gcd)
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   116
  
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   117
lemma semiring_Gcd:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64786
diff 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: 64785
diff changeset
   120
proof -
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   121
  show ?thesis
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   122
  proof
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   124
    proof (cases "\<exists>l. l \<noteq>  0 \<and> (\<forall>a\<in>A. a dvd l)")
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   125
      case False
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   126
      then have "Lcm A = 0"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   128
      with False show ?thesis
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   129
        by auto
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   130
    next
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   131
      case True
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   136
        apply (subst n_def)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   140
        done
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   142
        and "euclidean_size l = n" 
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   144
      {
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   147
          by (auto intro: gcd_greatest)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   149
          by simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   156
        proof -
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   158
            by simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   161
            by auto
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   163
            by (rule size_mult_mono)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   167
        qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   173
      }
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   177
        by auto
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   179
        by (simp add: local.Lcm_def Let_def n_def l_def)
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   180
      finally show ?thesis .
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   181
    qed
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   184
      by auto
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   192
      by (simp add: local.Gcd_def)
62442
26e4be6a680f More efficient Extended Euclidean Algorithm
Manuel Eberl <eberlm@in.tum.de>
parents: 62429
diff changeset
   193
  qed
26e4be6a680f More efficient Extended Euclidean Algorithm
Manuel Eberl <eberlm@in.tum.de>
parents: 62429
diff changeset
   194
qed
26e4be6a680f More efficient Extended Euclidean Algorithm
Manuel Eberl <eberlm@in.tum.de>
parents: 62429
diff changeset
   195
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64786
diff 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: 64785
diff changeset
   198
  by (fact semiring_Gcd)
60598
78ca5674c66a rings follow immediately their corresponding semirings
haftmann
parents: 60582
diff changeset
   199
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   200
subclass factorial_semiring
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   201
proof -
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64786
diff changeset
   203
     unit_factor normalize"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64850
diff changeset
   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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   208
      case (less x)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   209
      show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   211
        case False
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   213
        proof
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   217
          proof (elim disjE)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   218
            assume "is_unit p"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   221
          next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   222
            assume "x dvd p"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   225
          qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   226
        qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   229
      next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   230
        case True
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff 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: 64785
diff changeset
   236
        have normalized_factors_product:
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   237
          "{p. p dvd a * b \<and> normalize p = p} \<subseteq>
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   238
             (\<lambda>(x,y). normalize (x * y)) ` ({p. p dvd a \<and> normalize p = p} \<times> {p. p dvd b \<and> normalize p = p})"
69884
dec7cc38a5dc tuned proof;
wenzelm
parents: 69064
diff changeset
   239
          for a b
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   240
        proof safe
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   241
          fix p assume p: "p dvd a * b" "normalize p = p"
69884
dec7cc38a5dc tuned proof;
wenzelm
parents: 69064
diff changeset
   242
          from p(1) obtain x y where xy: "p = x * y" "x dvd a" "y dvd b"
dec7cc38a5dc tuned proof;
wenzelm
parents: 69064
diff changeset
   243
            by (rule dvd_productE)
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   244
          define x' y' where "x' = normalize x" and "y' = normalize y"
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   245
          have "p = normalize (x' * y')"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   246
            using p by (simp add: xy x'_def y'_def)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   247
          moreover have "x' dvd a \<and> normalize x' = x'" and "y' dvd b \<and> normalize y' = y'"
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   248
            using xy by (auto simp: x'_def y'_def)
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   249
          ultimately show "p \<in> (\<lambda>(x, y). normalize (x * y)) `
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   250
              ({p. p dvd a \<and> normalize p = p} \<times> {p. p dvd b \<and> normalize p = p})" by fast
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   251
        qed
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   252
        from x y have "\<not>is_unit z" by (auto simp: mult_unit_dvd_iff)
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   253
        have "?fctrs x \<subseteq> (\<lambda>(p,p'). normalize (p * p')) ` (?fctrs y \<times> ?fctrs z)"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   254
          by (subst x) (rule normalized_factors_product)
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   255
        moreover have "\<not>y * z dvd y * 1" "\<not>y * z dvd 1 * z"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   256
          by (subst dvd_times_left_cancel_iff dvd_times_right_cancel_iff; fact)+
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   257
        hence "finite ((\<lambda>(p,p'). normalize (p * p')) ` (?fctrs y \<times> ?fctrs z))"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   258
          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: 64785
diff changeset
   259
             (auto simp: x)
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   260
        ultimately show ?thesis by (rule finite_subset)
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   261
      qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   262
    qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   263
  next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   264
    fix p
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   265
    assume "irreducible p"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   266
    then show "prime_elem p"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   267
      by (rule irreducible_imp_prime_elem_gcd)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   268
  qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   269
qed
60598
78ca5674c66a rings follow immediately their corresponding semirings
haftmann
parents: 60582
diff changeset
   270
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   271
lemma Gcd_eucl_set [code]:
64850
fc9265882329 gcd/lcm on finite sets
haftmann
parents: 64848
diff changeset
   272
  "Gcd (set xs) = fold gcd xs 0"
fc9265882329 gcd/lcm on finite sets
haftmann
parents: 64848
diff changeset
   273
  by (fact Gcd_set_eq_fold)
60598
78ca5674c66a rings follow immediately their corresponding semirings
haftmann
parents: 60582
diff changeset
   274
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   275
lemma Lcm_eucl_set [code]:
64850
fc9265882329 gcd/lcm on finite sets
haftmann
parents: 64848
diff changeset
   276
  "Lcm (set xs) = fold lcm xs 1"
fc9265882329 gcd/lcm on finite sets
haftmann
parents: 64848
diff changeset
   277
  by (fact Lcm_set_eq_fold)
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   278
 
60598
78ca5674c66a rings follow immediately their corresponding semirings
haftmann
parents: 60582
diff changeset
   279
end
78ca5674c66a rings follow immediately their corresponding semirings
haftmann
parents: 60582
diff changeset
   280
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   281
hide_const (open) gcd lcm Gcd Lcm
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   282
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   283
lemma prime_elem_int_abs_iff [simp]:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   284
  fixes p :: int
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   285
  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: 64785
diff changeset
   286
  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: 64785
diff changeset
   287
  
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   288
lemma prime_elem_int_minus_iff [simp]:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   289
  fixes p :: int
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   290
  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: 64785
diff changeset
   291
  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: 64785
diff changeset
   292
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   293
lemma prime_int_iff:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   294
  fixes p :: int
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   295
  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: 64785
diff changeset
   296
  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: 64785
diff changeset
   297
  
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   298
  
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   299
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: 64785
diff changeset
   300
  
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
   301
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: 64785
diff changeset
   302
  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: 64785
diff changeset
   303
    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: 64785
diff changeset
   304
  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: 64785
diff changeset
   305
    and Lcm_eucl: "Euclidean_Algorithm.Lcm = GCD.Lcm"
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   306
begin
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   307
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   308
subclass semiring_gcd
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   309
  unfolding gcd_eucl [symmetric] lcm_eucl [symmetric]
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   310
  by (fact semiring_gcd)
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   311
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   312
subclass semiring_Gcd
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   313
  unfolding  gcd_eucl [symmetric] lcm_eucl [symmetric]
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   314
    Gcd_eucl [symmetric] Lcm_eucl [symmetric]
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   315
  by (fact semiring_Gcd)
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   316
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   317
subclass factorial_semiring_gcd
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   318
proof
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   319
  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: 64785
diff changeset
   320
    apply (rule sym)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   321
    apply (rule gcdI)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   322
       apply (fact gcd_lcm_factorial)+
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   323
    done
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   324
  then show "lcm a b = lcm_factorial a b" for a b
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   325
    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: 64785
diff changeset
   326
  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: 64785
diff changeset
   327
    apply (rule sym)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   328
    apply (rule GcdI)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   329
       apply (fact gcd_lcm_factorial)+
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   330
    done
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   331
  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: 64785
diff changeset
   332
    apply (rule sym)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   333
    apply (rule LcmI)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   334
       apply (fact gcd_lcm_factorial)+
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   335
    done
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   336
qed
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   337
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   338
lemma gcd_mod_right [simp]:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   339
  "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: 64785
diff changeset
   340
  unfolding gcd.commute [of a b]
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   341
  by (simp add: gcd_eucl [symmetric] local.gcd_mod)
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   342
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   343
lemma gcd_mod_left [simp]:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   344
  "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: 64785
diff changeset
   345
  by (drule gcd_mod_right [of _ a]) (simp add: gcd.commute)
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   346
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   347
lemma euclidean_size_gcd_le1 [simp]:
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   348
  assumes "a \<noteq> 0"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   349
  shows "euclidean_size (gcd a b) \<le> euclidean_size a"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   350
proof -
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   351
  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: 64785
diff changeset
   352
  with assms have "c \<noteq> 0"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   353
    by auto
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   354
  moreover from this
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   355
  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: 64785
diff changeset
   356
    by (rule size_mult_mono)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   357
  with A show ?thesis
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   358
    by simp
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   359
qed
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   360
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   361
lemma euclidean_size_gcd_le2 [simp]:
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   362
  "b \<noteq> 0 \<Longrightarrow> euclidean_size (gcd a b) \<le> euclidean_size b"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   363
  by (subst gcd.commute, rule euclidean_size_gcd_le1)
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   364
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   365
lemma euclidean_size_gcd_less1:
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   366
  assumes "a \<noteq> 0" and "\<not> a dvd b"
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   367
  shows "euclidean_size (gcd a b) < euclidean_size a"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   368
proof (rule ccontr)
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   369
  assume "\<not>euclidean_size (gcd a b) < euclidean_size a"
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   370
  with \<open>a \<noteq> 0\<close> have A: "euclidean_size (gcd a b) = euclidean_size a"
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   371
    by (intro le_antisym, simp_all)
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   372
  have "a dvd gcd a b"
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   373
    by (rule dvd_euclidean_size_eq_imp_dvd) (simp_all add: assms A)
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   374
  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: 64785
diff changeset
   375
  with \<open>\<not> a dvd b\<close> show False by contradiction
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   376
qed
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   377
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   378
lemma euclidean_size_gcd_less2:
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   379
  assumes "b \<noteq> 0" and "\<not> b dvd a"
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   380
  shows "euclidean_size (gcd a b) < euclidean_size b"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   381
  using assms by (subst gcd.commute, rule euclidean_size_gcd_less1)
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   382
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   383
lemma euclidean_size_lcm_le1: 
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   384
  assumes "a \<noteq> 0" and "b \<noteq> 0"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   385
  shows "euclidean_size a \<le> euclidean_size (lcm a b)"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   386
proof -
60690
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   387
  have "a dvd lcm a b" by (rule dvd_lcm1)
a9e45c9588c3 tuned facts
haftmann
parents: 60688
diff changeset
   388
  then obtain c where A: "lcm a b = a * c" ..
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62428
diff changeset
   389
  with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> have "c \<noteq> 0" by (auto simp: lcm_eq_0_iff)
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   390
  then show ?thesis by (subst A, intro size_mult_mono)
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   391
qed
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   392
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   393
lemma euclidean_size_lcm_le2:
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   394
  "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> euclidean_size b \<le> euclidean_size (lcm a b)"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   395
  using euclidean_size_lcm_le1 [of b a] by (simp add: ac_simps)
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   396
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   397
lemma euclidean_size_lcm_less1:
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   398
  assumes "b \<noteq> 0" and "\<not> b dvd a"
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   399
  shows "euclidean_size a < euclidean_size (lcm a b)"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   400
proof (rule ccontr)
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   401
  from assms have "a \<noteq> 0" by auto
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   402
  assume "\<not>euclidean_size a < euclidean_size (lcm a b)"
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 60517
diff changeset
   403
  with \<open>a \<noteq> 0\<close> and \<open>b \<noteq> 0\<close> have "euclidean_size (lcm a b) = euclidean_size a"
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   404
    by (intro le_antisym, simp, intro euclidean_size_lcm_le1)
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   405
  with assms have "lcm a b dvd a" 
62429
25271ff79171 Tuned Euclidean Rings/GCD rings
Manuel Eberl <eberlm@in.tum.de>
parents: 62428
diff changeset
   406
    by (rule_tac dvd_euclidean_size_eq_imp_dvd) (auto simp: lcm_eq_0_iff)
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   407
  hence "b dvd a" by (rule lcm_dvdD2)
60526
fad653acf58f isabelle update_cartouches;
wenzelm
parents: 60517
diff changeset
   408
  with \<open>\<not>b dvd a\<close> show False by contradiction
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   409
qed
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   410
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   411
lemma euclidean_size_lcm_less2:
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   412
  assumes "a \<noteq> 0" and "\<not> a dvd b"
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   413
  shows "euclidean_size b < euclidean_size (lcm a b)"
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   414
  using assms euclidean_size_lcm_less1 [of a b] by (simp add: ac_simps)
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   415
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   416
end
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   417
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   418
lemma factorial_euclidean_semiring_gcdI:
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
   419
  "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: 64785
diff changeset
   420
proof
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   421
  interpret semiring_Gcd 1 0 times
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   422
    Euclidean_Algorithm.gcd Euclidean_Algorithm.lcm
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   423
    Euclidean_Algorithm.Gcd Euclidean_Algorithm.Lcm
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64786
diff changeset
   424
    divide plus minus unit_factor normalize
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67443
diff changeset
   425
    rewrites "dvd.dvd (*) = Rings.dvd"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   426
    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: 64785
diff changeset
   427
  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: 64785
diff changeset
   428
  proof (rule ext)+
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   429
    fix a b :: 'a
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   430
    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: 64785
diff changeset
   431
    proof (induct a b rule: eucl_induct)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   432
      case zero
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   433
      then show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   434
        by simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   435
    next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   436
      case (mod a b)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   437
      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: 64785
diff changeset
   438
        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: 64785
diff changeset
   439
          by (simp add: div_mult_mod_eq)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   440
      ultimately show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   441
        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: 64785
diff changeset
   442
    qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   443
  qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   444
  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: 64785
diff changeset
   445
    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: 64785
diff changeset
   446
  show "Euclidean_Algorithm.lcm = (lcm :: 'a \<Rightarrow> _)"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   447
    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: 64785
diff changeset
   448
  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: 64785
diff changeset
   449
    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: 64785
diff changeset
   450
qed
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   451
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   452
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   453
subsection \<open>The extended euclidean algorithm\<close>
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   454
  
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   455
class euclidean_ring_gcd = euclidean_semiring_gcd + idom
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   456
begin
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   457
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   458
subclass euclidean_ring ..
60439
b765e08f8bc0 proper subclass instances for existing gcd (semi)rings
haftmann
parents: 60438
diff changeset
   459
subclass ring_gcd ..
63498
a3fe3250d05d Reformed factorial rings
eberlm <eberlm@in.tum.de>
parents: 63167
diff changeset
   460
subclass factorial_ring_gcd ..
60439
b765e08f8bc0 proper subclass instances for existing gcd (semi)rings
haftmann
parents: 60438
diff changeset
   461
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   462
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: 64785
diff changeset
   463
  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: 64785
diff changeset
   464
     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: 64785
diff changeset
   465
     else let q = r' div r
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   466
          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: 64785
diff changeset
   467
  by auto
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   468
termination
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   469
  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: 64785
diff changeset
   470
    (simp_all add: mod_size_less)
62442
26e4be6a680f More efficient Extended Euclidean Algorithm
Manuel Eberl <eberlm@in.tum.de>
parents: 62429
diff changeset
   471
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   472
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: 64785
diff changeset
   473
  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: 64785
diff changeset
   474
    
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   475
lemma
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   476
  assumes "gcd r' r = gcd a b"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff 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: 64785
diff changeset
   478
  assumes "s * a + t * b = r"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   479
  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: 64785
diff changeset
   480
  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: 64785
diff changeset
   481
    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: 64785
diff changeset
   482
proof -
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   483
  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: 64785
diff changeset
   484
    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: 64785
diff changeset
   485
    using assms(1-3)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   486
  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: 64785
diff changeset
   487
    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: 64785
diff changeset
   488
    show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   489
    proof (cases "r = 0")
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   490
      case True
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   491
      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: 64785
diff changeset
   492
               ((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: 64785
diff changeset
   493
        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: 64785
diff changeset
   494
      also have "?P \<dots>"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   495
      proof safe
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   496
        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: 64785
diff changeset
   497
                (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: 64785
diff changeset
   498
          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: 64785
diff changeset
   499
        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: 64785
diff changeset
   500
        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: 64785
diff changeset
   501
        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: 64785
diff changeset
   502
      next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   503
        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: 64785
diff changeset
   504
          by simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   505
      qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   506
      finally show ?thesis .
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   507
    next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   508
      case False
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   509
      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: 64785
diff changeset
   510
             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: 64785
diff changeset
   511
        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: 64785
diff changeset
   512
      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: 64785
diff changeset
   513
      proof (intro "1.IH")
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   514
        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: 64785
diff changeset
   515
              (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: 64785
diff 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: 64785
diff changeset
   517
        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: 64785
diff changeset
   518
        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: 64785
diff changeset
   519
          by (simp add: algebra_simps)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   520
        finally show "(s' - r' div r * s) * a + (t' - r' div r * t) * b = r' mod r" .
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
   521
      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: 64785
diff changeset
   522
      finally show ?thesis .
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   523
    qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   524
  qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   525
  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: 64785
diff changeset
   526
    by simp_all
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   527
qed
60572
718b1ba06429 streamlined definitions and primitive lemma of euclidean algorithm, including code generation
haftmann
parents: 60571
diff changeset
   528
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   529
declare euclid_ext_aux.simps [simp del]
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   530
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   531
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: 64785
diff changeset
   532
  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: 64785
diff changeset
   533
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   534
lemma bezout_coefficients_0: 
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   535
  "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: 64785
diff changeset
   536
  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: 64785
diff changeset
   537
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   538
lemma bezout_coefficients_left_0: 
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   539
  "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: 64785
diff changeset
   540
  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: 64785
diff changeset
   541
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   542
lemma bezout_coefficients:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   543
  assumes "bezout_coefficients a b = (x, y)"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   544
  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: 64785
diff changeset
   545
  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: 64785
diff changeset
   546
    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: 64785
diff changeset
   547
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   548
lemma bezout_coefficients_fst_snd:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   549
  "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: 64785
diff changeset
   550
  by (rule bezout_coefficients) simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   551
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   552
lemma euclid_ext_eq [simp]:
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   553
  "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: 64785
diff changeset
   554
proof
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   555
  show "fst ?p = fst ?q"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   556
    by (simp add: bezout_coefficients_def)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   557
  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: 64785
diff changeset
   558
    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: 64785
diff changeset
   559
      (simp_all add: prod_eq_iff)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   560
  then show "snd ?p = snd ?q"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   561
    by simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   562
qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   563
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   564
declare euclid_ext_eq [symmetric, code_unfold]
60572
718b1ba06429 streamlined definitions and primitive lemma of euclidean algorithm, including code generation
haftmann
parents: 60571
diff changeset
   565
718b1ba06429 streamlined definitions and primitive lemma of euclidean algorithm, including code generation
haftmann
parents: 60571
diff changeset
   566
end
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   567
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   568
class normalization_euclidean_semiring_multiplicative =
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   569
  normalization_euclidean_semiring + normalization_semidom_multiplicative
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   570
begin
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   571
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   572
subclass factorial_semiring_multiplicative ..
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   573
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   574
end
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   575
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   576
class field_gcd =
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   577
  field + unique_euclidean_ring + euclidean_ring_gcd + normalization_semidom_multiplicative
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   578
begin
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   579
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   580
subclass normalization_euclidean_semiring_multiplicative ..
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   581
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   582
subclass normalization_euclidean_semiring ..
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   583
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   584
subclass semiring_gcd_mult_normalize ..
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   585
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   586
end
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   587
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   588
60572
718b1ba06429 streamlined definitions and primitive lemma of euclidean algorithm, including code generation
haftmann
parents: 60571
diff changeset
   589
subsection \<open>Typical instances\<close>
58023
62826b36ac5e generic euclidean algorithm (due to Manuel Eberl)
haftmann
parents:
diff changeset
   590
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
   591
instance nat :: normalization_euclidean_semiring ..
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
   592
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   593
instance nat :: euclidean_semiring_gcd
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   594
proof
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   595
  interpret semiring_Gcd 1 0 times
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   596
    "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   597
    "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm"
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64786
diff changeset
   598
    divide plus minus unit_factor normalize
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67443
diff changeset
   599
    rewrites "dvd.dvd (*) = Rings.dvd"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   600
    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: 64785
diff changeset
   601
  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: 64785
diff changeset
   602
  proof (rule ext)+
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   603
    fix m n :: nat
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   604
    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: 64785
diff changeset
   605
    proof (induct m n rule: eucl_induct)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   606
      case zero
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   607
      then show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   608
        by simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   609
    next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   610
      case (mod m n)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   611
      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: 64785
diff changeset
   612
        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: 64785
diff changeset
   613
      with mod show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   614
        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: 64785
diff changeset
   615
    qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   616
  qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   617
  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: 64785
diff changeset
   618
    by (auto intro!: ext Lcm_eqI)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   619
  show "(Euclidean_Algorithm.lcm :: nat \<Rightarrow> _) = lcm"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   620
    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: 64785
diff changeset
   621
  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: 64785
diff changeset
   622
    by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   623
qed
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   624
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   625
instance nat :: normalization_euclidean_semiring_multiplicative ..
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   626
67167
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 66817
diff changeset
   627
lemma prime_factorization_Suc_0 [simp]: "prime_factorization (Suc 0) = {#}"
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 66817
diff changeset
   628
  unfolding One_nat_def [symmetric] using prime_factorization_1 .
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 66817
diff changeset
   629
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
   630
instance int :: normalization_euclidean_semiring ..
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 65435
diff changeset
   631
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   632
instance int :: euclidean_ring_gcd
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   633
proof
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   634
  interpret semiring_Gcd 1 0 times
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   635
    "Euclidean_Algorithm.gcd" "Euclidean_Algorithm.lcm"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   636
    "Euclidean_Algorithm.Gcd" "Euclidean_Algorithm.Lcm"
64848
c50db2128048 slightly generalized type class hierarchy concerning unit factors, to allow for lean polynomial normalization
haftmann
parents: 64786
diff changeset
   637
    divide plus minus unit_factor normalize
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67443
diff changeset
   638
    rewrites "dvd.dvd (*) = Rings.dvd"
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   639
    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: 64785
diff changeset
   640
  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: 64785
diff changeset
   641
  proof (rule ext)+
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   642
    fix k l :: int
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   643
    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: 64785
diff changeset
   644
    proof (induct k l rule: eucl_induct)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   645
      case zero
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   646
      then show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   647
        by simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   648
    next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   649
      case (mod k l)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   650
      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: 64785
diff changeset
   651
      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: 64785
diff changeset
   652
        case less
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   653
        then show ?thesis
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   654
          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: 64785
diff changeset
   655
      next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   656
        case equal
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   657
        with mod show ?thesis
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   658
          by simp
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   659
      next
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   660
        case greater
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   661
        then show ?thesis
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   662
          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: 64785
diff changeset
   663
      qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   664
      with mod show ?case
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   665
        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: 64785
diff changeset
   666
    qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   667
  qed
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   668
  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: 64785
diff changeset
   669
    by (auto intro!: ext Lcm_eqI)
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   670
  show "(Euclidean_Algorithm.lcm :: int \<Rightarrow> _) = lcm"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64785
diff changeset
   671
    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: 64785
diff changeset
   672
  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: 64785
diff changeset
   673
    by (simp add: fun_eq_iff Euclidean_Algorithm.Gcd_def semiring_Gcd_class.Gcd_Lcm)
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   674
qed
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62353
diff changeset
   675
71398
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   676
instance int :: normalization_euclidean_semiring_multiplicative ..
e0237f2eb49d Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents: 69884
diff changeset
   677
80084
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   678
lemma (in idom) prime_CHAR_semidom:
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   679
  assumes "CHAR('a) > 0"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   680
  shows   "prime CHAR('a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   681
proof -
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   682
  have False if ab: "a \<noteq> 1" "b \<noteq> 1" "CHAR('a) = a * b" for a b
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   683
  proof -
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   684
    from assms ab have "a > 0" "b > 0"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   685
      by (auto intro!: Nat.gr0I)    
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   686
    have "of_nat (a * b) = (0 :: 'a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   687
      using ab by (metis of_nat_CHAR)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   688
    also have "of_nat (a * b) = (of_nat a :: 'a) * of_nat b"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   689
      by simp
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   690
    finally have "of_nat a * of_nat b = (0 :: 'a)" .
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   691
    moreover have "of_nat a * of_nat b \<noteq> (0 :: 'a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   692
      using ab \<open>a > 0\<close> \<open>b > 0\<close>
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   693
      by (intro no_zero_divisors) (auto simp: of_nat_eq_0_iff_char_dvd)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   694
    ultimately show False
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   695
      by contradiction
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   696
  qed
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   697
  moreover have "CHAR('a) > 1"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   698
    using assms CHAR_not_1' by linarith
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   699
  ultimately have "prime_elem CHAR('a)"
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   700
    by (intro irreducible_imp_prime_elem) (auto simp: Factorial_Ring.irreducible_def)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   701
  thus ?thesis
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   702
    by (auto simp: prime_def)
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   703
qed
173548e4d5d0 moved over material from the AFP to HOL, HOL-Computational_Algebra, and HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 71398
diff changeset
   704
63924
f91766530e13 more generic algebraic lemmas
haftmann
parents: 63633
diff changeset
   705
end