src/ZF/ex/Primes.thy
author paulson
Wed, 27 Apr 2005 16:41:03 +0200
changeset 15863 78db9506cc78
parent 13339 0f89104dd377
child 16417 9bc16273c2d4
permissions -rw-r--r--
minor tidying
paulson@1792
     1
(*  Title:      ZF/ex/Primes.thy
paulson@1792
     2
    ID:         $Id$
paulson@1792
     3
    Author:     Christophe Tabacznyj and Lawrence C Paulson
paulson@1792
     4
    Copyright   1996  University of Cambridge
paulson@15863
     5
*)
paulson@1792
     6
paulson@15863
     7
header{*The Divides Relation and Euclid's algorithm for the GCD*}
paulson@1792
     8
paulson@12867
     9
theory Primes = Main:
paulson@12867
    10
constdefs
paulson@12867
    11
  divides :: "[i,i]=>o"              (infixl "dvd" 50) 
paulson@12867
    12
    "m dvd n == m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
paulson@1792
    13
paulson@15863
    14
  is_gcd  :: "[i,i,i]=>o"     --{*definition of great common divisor*}
paulson@12867
    15
    "is_gcd(p,m,n) == ((p dvd m) & (p dvd n))   &
paulson@12867
    16
                       (\<forall>d\<in>nat. (d dvd m) & (d dvd n) --> d dvd p)"
paulson@1792
    17
paulson@15863
    18
  gcd     :: "[i,i]=>i"       --{*Euclid's algorithm for the gcd*}
paulson@12867
    19
    "gcd(m,n) == transrec(natify(n),
paulson@12867
    20
			%n f. \<lambda>m \<in> nat.
paulson@11382
    21
			        if n=0 then m else f`(m mod n)`n) ` natify(m)"
paulson@1792
    22
paulson@15863
    23
  coprime :: "[i,i]=>o"       --{*the coprime relation*}
paulson@11382
    24
    "coprime(m,n) == gcd(m,n) = 1"
paulson@11382
    25
  
paulson@15863
    26
  prime   :: i                --{*the set of prime numbers*}
paulson@12867
    27
   "prime == {p \<in> nat. 1<p & (\<forall>m \<in> nat. m dvd p --> m=1 | m=p)}"
paulson@12867
    28
paulson@15863
    29
paulson@15863
    30
subsection{*The Divides Relation*}
paulson@12867
    31
paulson@12867
    32
lemma dvdD: "m dvd n ==> m \<in> nat & n \<in> nat & (\<exists>k \<in> nat. n = m#*k)"
paulson@13339
    33
by (unfold divides_def, assumption)
paulson@12867
    34
paulson@12867
    35
lemma dvdE:
paulson@12867
    36
     "[|m dvd n;  !!k. [|m \<in> nat; n \<in> nat; k \<in> nat; n = m#*k|] ==> P|] ==> P"
paulson@12867
    37
by (blast dest!: dvdD)
paulson@12867
    38
paulson@12867
    39
lemmas dvd_imp_nat1 = dvdD [THEN conjunct1, standard]
paulson@12867
    40
lemmas dvd_imp_nat2 = dvdD [THEN conjunct2, THEN conjunct1, standard]
paulson@12867
    41
paulson@12867
    42
paulson@12867
    43
lemma dvd_0_right [simp]: "m \<in> nat ==> m dvd 0"
paulson@15863
    44
apply (simp add: divides_def)
paulson@12867
    45
apply (fast intro: nat_0I mult_0_right [symmetric])
paulson@12867
    46
done
paulson@12867
    47
paulson@12867
    48
lemma dvd_0_left: "0 dvd m ==> m = 0"
paulson@15863
    49
by (simp add: divides_def)
paulson@12867
    50
paulson@12867
    51
lemma dvd_refl [simp]: "m \<in> nat ==> m dvd m"
paulson@15863
    52
apply (simp add: divides_def)
paulson@12867
    53
apply (fast intro: nat_1I mult_1_right [symmetric])
paulson@12867
    54
done
paulson@12867
    55
paulson@12867
    56
lemma dvd_trans: "[| m dvd n; n dvd p |] ==> m dvd p"
paulson@15863
    57
by (auto simp add: divides_def intro: mult_assoc mult_type)
paulson@12867
    58
paulson@12867
    59
lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m=n"
paulson@15863
    60
apply (simp add: divides_def)
paulson@12867
    61
apply (force dest: mult_eq_self_implies_10
paulson@12867
    62
             simp add: mult_assoc mult_eq_1_iff)
paulson@12867
    63
done
paulson@12867
    64
paulson@12867
    65
lemma dvd_mult_left: "[|(i#*j) dvd k; i \<in> nat|] ==> i dvd k"
paulson@15863
    66
by (auto simp add: divides_def mult_assoc)
paulson@12867
    67
paulson@12867
    68
lemma dvd_mult_right: "[|(i#*j) dvd k; j \<in> nat|] ==> j dvd k"
paulson@15863
    69
apply (simp add: divides_def, clarify)
paulson@12867
    70
apply (rule_tac x = "i#*k" in bexI)
paulson@12867
    71
apply (simp add: mult_ac)
paulson@12867
    72
apply (rule mult_type)
paulson@12867
    73
done
paulson@12867
    74
paulson@12867
    75
paulson@15863
    76
subsection{*Euclid's Algorithm for the GCD*}
paulson@12867
    77
paulson@12867
    78
lemma gcd_0 [simp]: "gcd(m,0) = natify(m)"
paulson@15863
    79
apply (simp add: gcd_def)
paulson@13339
    80
apply (subst transrec, simp)
paulson@12867
    81
done
paulson@12867
    82
paulson@12867
    83
lemma gcd_natify1 [simp]: "gcd(natify(m),n) = gcd(m,n)"
paulson@12867
    84
by (simp add: gcd_def)
paulson@12867
    85
paulson@12867
    86
lemma gcd_natify2 [simp]: "gcd(m, natify(n)) = gcd(m,n)"
paulson@12867
    87
by (simp add: gcd_def)
paulson@12867
    88
paulson@12867
    89
lemma gcd_non_0_raw: 
paulson@12867
    90
    "[| 0<n;  n \<in> nat |] ==> gcd(m,n) = gcd(n, m mod n)"
paulson@15863
    91
apply (simp add: gcd_def)
paulson@12867
    92
apply (rule_tac P = "%z. ?left (z) = ?right" in transrec [THEN ssubst])
paulson@12867
    93
apply (simp add: ltD [THEN mem_imp_not_eq, THEN not_sym] 
paulson@12867
    94
                 mod_less_divisor [THEN ltD])
paulson@12867
    95
done
paulson@12867
    96
paulson@12867
    97
lemma gcd_non_0: "0 < natify(n) ==> gcd(m,n) = gcd(n, m mod n)"
paulson@13339
    98
apply (cut_tac m = m and n = "natify (n) " in gcd_non_0_raw)
paulson@12867
    99
apply auto
paulson@12867
   100
done
paulson@12867
   101
paulson@12867
   102
lemma gcd_1 [simp]: "gcd(m,1) = 1"
paulson@12867
   103
by (simp (no_asm_simp) add: gcd_non_0)
paulson@12867
   104
paulson@12867
   105
lemma dvd_add: "[| k dvd a; k dvd b |] ==> k dvd (a #+ b)"
paulson@15863
   106
apply (simp add: divides_def)
paulson@12867
   107
apply (fast intro: add_mult_distrib_left [symmetric] add_type)
paulson@12867
   108
done
paulson@12867
   109
paulson@12867
   110
lemma dvd_mult: "k dvd n ==> k dvd (m #* n)"
paulson@15863
   111
apply (simp add: divides_def)
paulson@12867
   112
apply (fast intro: mult_left_commute mult_type)
paulson@12867
   113
done
paulson@12867
   114
paulson@12867
   115
lemma dvd_mult2: "k dvd m ==> k dvd (m #* n)"
paulson@12867
   116
apply (subst mult_commute)
paulson@12867
   117
apply (blast intro: dvd_mult)
paulson@12867
   118
done
paulson@12867
   119
paulson@12867
   120
(* k dvd (m*k) *)
paulson@12867
   121
lemmas dvdI1 [simp] = dvd_refl [THEN dvd_mult, standard]
paulson@12867
   122
lemmas dvdI2 [simp] = dvd_refl [THEN dvd_mult2, standard]
paulson@12867
   123
paulson@12867
   124
lemma dvd_mod_imp_dvd_raw:
paulson@12867
   125
     "[| a \<in> nat; b \<in> nat; k dvd b; k dvd (a mod b) |] ==> k dvd a"
paulson@15863
   126
apply (case_tac "b=0") 
paulson@13259
   127
 apply (simp add: DIVISION_BY_ZERO_MOD)
paulson@12867
   128
apply (blast intro: mod_div_equality [THEN subst]
paulson@12867
   129
             elim: dvdE 
paulson@12867
   130
             intro!: dvd_add dvd_mult mult_type mod_type div_type)
paulson@12867
   131
done
paulson@12867
   132
paulson@12867
   133
lemma dvd_mod_imp_dvd: "[| k dvd (a mod b); k dvd b; a \<in> nat |] ==> k dvd a"
paulson@13259
   134
apply (cut_tac b = "natify (b)" in dvd_mod_imp_dvd_raw)
paulson@12867
   135
apply auto
paulson@12867
   136
apply (simp add: divides_def)
paulson@12867
   137
done
paulson@12867
   138
paulson@12867
   139
(*Imitating TFL*)
paulson@12867
   140
lemma gcd_induct_lemma [rule_format (no_asm)]: "[| n \<in> nat;  
paulson@12867
   141
         \<forall>m \<in> nat. P(m,0);  
paulson@12867
   142
         \<forall>m \<in> nat. \<forall>n \<in> nat. 0<n --> P(n, m mod n) --> P(m,n) |]  
paulson@12867
   143
      ==> \<forall>m \<in> nat. P (m,n)"
paulson@13339
   144
apply (erule_tac i = n in complete_induct)
paulson@12867
   145
apply (case_tac "x=0")
paulson@12867
   146
apply (simp (no_asm_simp))
paulson@12867
   147
apply clarify
paulson@13339
   148
apply (drule_tac x1 = m and x = x in bspec [THEN bspec])
paulson@12867
   149
apply (simp_all add: Ord_0_lt_iff)
paulson@12867
   150
apply (blast intro: mod_less_divisor [THEN ltD])
paulson@12867
   151
done
paulson@12867
   152
paulson@12867
   153
lemma gcd_induct: "!!P. [| m \<in> nat; n \<in> nat;  
paulson@12867
   154
         !!m. m \<in> nat ==> P(m,0);  
paulson@12867
   155
         !!m n. [|m \<in> nat; n \<in> nat; 0<n; P(n, m mod n)|] ==> P(m,n) |]  
paulson@12867
   156
      ==> P (m,n)"
paulson@12867
   157
by (blast intro: gcd_induct_lemma)
paulson@12867
   158
paulson@12867
   159
paulson@15863
   160
subsection{*Basic Properties of @{term gcd}*}
paulson@12867
   161
paulson@15863
   162
text{*type of gcd*}
paulson@12867
   163
lemma gcd_type [simp,TC]: "gcd(m, n) \<in> nat"
paulson@13339
   164
apply (subgoal_tac "gcd (natify (m), natify (n)) \<in> nat")
paulson@12867
   165
apply simp
paulson@13259
   166
apply (rule_tac m = "natify (m)" and n = "natify (n)" in gcd_induct)
paulson@12867
   167
apply auto
paulson@12867
   168
apply (simp add: gcd_non_0)
paulson@12867
   169
done
paulson@12867
   170
paulson@12867
   171
paulson@15863
   172
text{* Property 1: gcd(a,b) divides a and b *}
paulson@12867
   173
paulson@12867
   174
lemma gcd_dvd_both:
paulson@12867
   175
     "[| m \<in> nat; n \<in> nat |] ==> gcd (m, n) dvd m & gcd (m, n) dvd n"
paulson@13339
   176
apply (rule_tac m = m and n = n in gcd_induct)
paulson@12867
   177
apply (simp_all add: gcd_non_0)
paulson@12867
   178
apply (blast intro: dvd_mod_imp_dvd_raw nat_into_Ord [THEN Ord_0_lt])
paulson@12867
   179
done
paulson@12867
   180
paulson@12867
   181
lemma gcd_dvd1 [simp]: "m \<in> nat ==> gcd(m,n) dvd m"
paulson@13259
   182
apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_dvd_both)
paulson@12867
   183
apply auto
paulson@12867
   184
done
paulson@12867
   185
paulson@12867
   186
lemma gcd_dvd2 [simp]: "n \<in> nat ==> gcd(m,n) dvd n"
paulson@13259
   187
apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_dvd_both)
paulson@12867
   188
apply auto
paulson@12867
   189
done
paulson@12867
   190
paulson@15863
   191
text{* if f divides a and b then f divides gcd(a,b) *}
paulson@12867
   192
paulson@12867
   193
lemma dvd_mod: "[| f dvd a; f dvd b |] ==> f dvd (a mod b)"
paulson@15863
   194
apply (simp add: divides_def)
paulson@13259
   195
apply (case_tac "b=0")
paulson@13339
   196
 apply (simp add: DIVISION_BY_ZERO_MOD, auto)
paulson@12867
   197
apply (blast intro: mod_mult_distrib2 [symmetric])
paulson@12867
   198
done
paulson@12867
   199
paulson@15863
   200
text{* Property 2: for all a,b,f naturals, 
paulson@15863
   201
               if f divides a and f divides b then f divides gcd(a,b)*}
paulson@12867
   202
paulson@12867
   203
lemma gcd_greatest_raw [rule_format]:
paulson@12867
   204
     "[| m \<in> nat; n \<in> nat; f \<in> nat |]    
paulson@12867
   205
      ==> (f dvd m) --> (f dvd n) --> f dvd gcd(m,n)"
paulson@13339
   206
apply (rule_tac m = m and n = n in gcd_induct)
paulson@12867
   207
apply (simp_all add: gcd_non_0 dvd_mod)
paulson@12867
   208
done
paulson@12867
   209
paulson@12867
   210
lemma gcd_greatest: "[| f dvd m;  f dvd n;  f \<in> nat |] ==> f dvd gcd(m,n)"
paulson@12867
   211
apply (rule gcd_greatest_raw)
paulson@12867
   212
apply (auto simp add: divides_def)
paulson@12867
   213
done
paulson@12867
   214
paulson@12867
   215
lemma gcd_greatest_iff [simp]: "[| k \<in> nat; m \<in> nat; n \<in> nat |]  
paulson@12867
   216
      ==> (k dvd gcd (m, n)) <-> (k dvd m & k dvd n)"
paulson@12867
   217
by (blast intro!: gcd_greatest gcd_dvd1 gcd_dvd2 intro: dvd_trans)
paulson@12867
   218
paulson@12867
   219
paulson@15863
   220
subsection{*The Greatest Common Divisor*}
paulson@15863
   221
paulson@15863
   222
text{*The GCD exists and function gcd computes it.*}
paulson@12867
   223
paulson@12867
   224
lemma is_gcd: "[| m \<in> nat; n \<in> nat |] ==> is_gcd(gcd(m,n), m, n)"
paulson@12867
   225
by (simp add: is_gcd_def)
paulson@12867
   226
paulson@15863
   227
text{*The GCD is unique*}
paulson@12867
   228
paulson@12867
   229
lemma is_gcd_unique: "[|is_gcd(m,a,b); is_gcd(n,a,b); m\<in>nat; n\<in>nat|] ==> m=n"
paulson@15863
   230
apply (simp add: is_gcd_def)
paulson@12867
   231
apply (blast intro: dvd_anti_sym)
paulson@12867
   232
done
paulson@12867
   233
paulson@12867
   234
lemma is_gcd_commute: "is_gcd(k,m,n) <-> is_gcd(k,n,m)"
paulson@15863
   235
by (simp add: is_gcd_def, blast)
paulson@12867
   236
paulson@12867
   237
lemma gcd_commute_raw: "[| m \<in> nat; n \<in> nat |] ==> gcd(m,n) = gcd(n,m)"
paulson@12867
   238
apply (rule is_gcd_unique)
paulson@12867
   239
apply (rule is_gcd)
paulson@12867
   240
apply (rule_tac [3] is_gcd_commute [THEN iffD1])
paulson@13339
   241
apply (rule_tac [3] is_gcd, auto)
paulson@12867
   242
done
paulson@12867
   243
paulson@12867
   244
lemma gcd_commute: "gcd(m,n) = gcd(n,m)"
paulson@13259
   245
apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_commute_raw)
paulson@12867
   246
apply auto
paulson@12867
   247
done
paulson@12867
   248
paulson@12867
   249
lemma gcd_assoc_raw: "[| k \<in> nat; m \<in> nat; n \<in> nat |]  
paulson@12867
   250
      ==> gcd (gcd (k, m), n) = gcd (k, gcd (m, n))"
paulson@12867
   251
apply (rule is_gcd_unique)
paulson@12867
   252
apply (rule is_gcd)
paulson@12867
   253
apply (simp_all add: is_gcd_def)
paulson@12867
   254
apply (blast intro: gcd_dvd1 gcd_dvd2 gcd_type intro: dvd_trans)
paulson@12867
   255
done
paulson@12867
   256
paulson@12867
   257
lemma gcd_assoc: "gcd (gcd (k, m), n) = gcd (k, gcd (m, n))"
paulson@13259
   258
apply (cut_tac k = "natify (k)" and m = "natify (m)" and n = "natify (n) " 
paulson@12867
   259
       in gcd_assoc_raw)
paulson@12867
   260
apply auto
paulson@12867
   261
done
paulson@12867
   262
paulson@12867
   263
lemma gcd_0_left [simp]: "gcd (0, m) = natify(m)"
paulson@12867
   264
by (simp add: gcd_commute [of 0])
paulson@12867
   265
paulson@12867
   266
lemma gcd_1_left [simp]: "gcd (1, m) = 1"
paulson@12867
   267
by (simp add: gcd_commute [of 1])
paulson@12867
   268
paulson@12867
   269
paulson@15863
   270
subsection{*Addition laws*}
paulson@15863
   271
paulson@15863
   272
lemma gcd_add1 [simp]: "gcd (m #+ n, n) = gcd (m, n)"
paulson@15863
   273
apply (subgoal_tac "gcd (m #+ natify (n), natify (n)) = gcd (m, natify (n))")
paulson@15863
   274
apply simp
paulson@15863
   275
apply (case_tac "natify (n) = 0")
paulson@15863
   276
apply (auto simp add: Ord_0_lt_iff gcd_non_0)
paulson@15863
   277
done
paulson@15863
   278
paulson@15863
   279
lemma gcd_add2 [simp]: "gcd (m, m #+ n) = gcd (m, n)"
paulson@15863
   280
apply (rule gcd_commute [THEN trans])
paulson@15863
   281
apply (subst add_commute, simp)
paulson@15863
   282
apply (rule gcd_commute)
paulson@15863
   283
done
paulson@15863
   284
paulson@15863
   285
lemma gcd_add2' [simp]: "gcd (m, n #+ m) = gcd (m, n)"
paulson@15863
   286
by (subst add_commute, rule gcd_add2)
paulson@15863
   287
paulson@15863
   288
lemma gcd_add_mult_raw: "k \<in> nat ==> gcd (m, k #* m #+ n) = gcd (m, n)"
paulson@15863
   289
apply (erule nat_induct)
paulson@15863
   290
apply (auto simp add: gcd_add2 add_assoc)
paulson@15863
   291
done
paulson@15863
   292
paulson@15863
   293
lemma gcd_add_mult: "gcd (m, k #* m #+ n) = gcd (m, n)"
paulson@15863
   294
apply (cut_tac k = "natify (k)" in gcd_add_mult_raw)
paulson@15863
   295
apply auto
paulson@15863
   296
done
paulson@15863
   297
paulson@15863
   298
paulson@15863
   299
subsection{* Multiplication Laws*}
paulson@12867
   300
paulson@12867
   301
lemma gcd_mult_distrib2_raw:
paulson@12867
   302
     "[| k \<in> nat; m \<in> nat; n \<in> nat |]  
paulson@12867
   303
      ==> k #* gcd (m, n) = gcd (k #* m, k #* n)"
paulson@13339
   304
apply (erule_tac m = m and n = n in gcd_induct, assumption)
paulson@12867
   305
apply simp
paulson@13339
   306
apply (case_tac "k = 0", simp)
paulson@12867
   307
apply (simp add: mod_geq gcd_non_0 mod_mult_distrib2 Ord_0_lt_iff)
paulson@12867
   308
done
paulson@12867
   309
paulson@12867
   310
lemma gcd_mult_distrib2: "k #* gcd (m, n) = gcd (k #* m, k #* n)"
paulson@13259
   311
apply (cut_tac k = "natify (k)" and m = "natify (m)" and n = "natify (n) " 
paulson@12867
   312
       in gcd_mult_distrib2_raw)
paulson@12867
   313
apply auto
paulson@12867
   314
done
paulson@12867
   315
paulson@12867
   316
lemma gcd_mult [simp]: "gcd (k, k #* n) = natify(k)"
paulson@13339
   317
by (cut_tac k = k and m = 1 and n = n in gcd_mult_distrib2, auto)
paulson@12867
   318
paulson@12867
   319
lemma gcd_self [simp]: "gcd (k, k) = natify(k)"
paulson@13339
   320
by (cut_tac k = k and n = 1 in gcd_mult, auto)
paulson@12867
   321
paulson@12867
   322
lemma relprime_dvd_mult:
paulson@12867
   323
     "[| gcd (k,n) = 1;  k dvd (m #* n);  m \<in> nat |] ==> k dvd m"
paulson@13339
   324
apply (cut_tac k = m and m = k and n = n in gcd_mult_distrib2, auto)
paulson@13339
   325
apply (erule_tac b = m in ssubst)
paulson@12867
   326
apply (simp add: dvd_imp_nat1)
paulson@12867
   327
done
paulson@12867
   328
paulson@12867
   329
lemma relprime_dvd_mult_iff:
paulson@12867
   330
     "[| gcd (k,n) = 1;  m \<in> nat |] ==> k dvd (m #* n) <-> k dvd m"
paulson@12867
   331
by (blast intro: dvdI2 relprime_dvd_mult dvd_trans)
paulson@12867
   332
paulson@12867
   333
lemma prime_imp_relprime: 
paulson@12867
   334
     "[| p \<in> prime;  ~ (p dvd n);  n \<in> nat |] ==> gcd (p, n) = 1"
paulson@15863
   335
apply (simp add: prime_def, clarify)
paulson@13259
   336
apply (drule_tac x = "gcd (p,n)" in bspec)
paulson@12867
   337
apply auto
paulson@13339
   338
apply (cut_tac m = p and n = n in gcd_dvd2, auto)
paulson@12867
   339
done
paulson@12867
   340
paulson@12867
   341
lemma prime_into_nat: "p \<in> prime ==> p \<in> nat"
paulson@15863
   342
by (simp add: prime_def)
paulson@12867
   343
paulson@12867
   344
lemma prime_nonzero: "p \<in> prime \<Longrightarrow> p\<noteq>0"
paulson@15863
   345
by (auto simp add: prime_def)
paulson@12867
   346
paulson@12867
   347
paulson@15863
   348
text{*This theorem leads immediately to a proof of the uniqueness of
paulson@15863
   349
  factorization.  If @{term p} divides a product of primes then it is
paulson@15863
   350
  one of those primes.*}
paulson@12867
   351
paulson@12867
   352
lemma prime_dvd_mult:
paulson@12867
   353
     "[|p dvd m #* n; p \<in> prime; m \<in> nat; n \<in> nat |] ==> p dvd m \<or> p dvd n"
paulson@12867
   354
by (blast intro: relprime_dvd_mult prime_imp_relprime prime_into_nat)
paulson@12867
   355
paulson@12867
   356
paulson@12867
   357
lemma gcd_mult_cancel_raw:
paulson@12867
   358
     "[|gcd (k,n) = 1; m \<in> nat; n \<in> nat|] ==> gcd (k #* m, n) = gcd (m, n)"
paulson@12867
   359
apply (rule dvd_anti_sym)
paulson@12867
   360
 apply (rule gcd_greatest)
paulson@12867
   361
  apply (rule relprime_dvd_mult [of _ k])
paulson@12867
   362
apply (simp add: gcd_assoc)
paulson@12867
   363
apply (simp add: gcd_commute)
paulson@12867
   364
apply (simp_all add: mult_commute)
paulson@12867
   365
apply (blast intro: dvdI1 gcd_dvd1 dvd_trans)
paulson@12867
   366
done
paulson@12867
   367
paulson@12867
   368
lemma gcd_mult_cancel: "gcd (k,n) = 1 ==> gcd (k #* m, n) = gcd (m, n)"
paulson@13259
   369
apply (cut_tac m = "natify (m)" and n = "natify (n)" in gcd_mult_cancel_raw)
paulson@12867
   370
apply auto
paulson@12867
   371
done
paulson@12867
   372
paulson@12867
   373
paulson@15863
   374
subsection{*The Square Root of a Prime is Irrational: Key Lemma*}
paulson@12867
   375
paulson@12867
   376
lemma prime_dvd_other_side:
paulson@12867
   377
     "\<lbrakk>n#*n = p#*(k#*k); p \<in> prime; n \<in> nat\<rbrakk> \<Longrightarrow> p dvd n"
paulson@12867
   378
apply (subgoal_tac "p dvd n#*n")
paulson@12867
   379
 apply (blast dest: prime_dvd_mult)
paulson@12867
   380
apply (rule_tac j = "k#*k" in dvd_mult_left)
paulson@12867
   381
 apply (auto simp add: prime_def)
paulson@12867
   382
done
paulson@12867
   383
paulson@12867
   384
lemma reduction:
paulson@12867
   385
     "\<lbrakk>k#*k = p#*(j#*j); p \<in> prime; 0 < k; j \<in> nat; k \<in> nat\<rbrakk>  
paulson@12867
   386
      \<Longrightarrow> k < p#*j & 0 < j"
paulson@12867
   387
apply (rule ccontr)
paulson@12867
   388
apply (simp add: not_lt_iff_le prime_into_nat)
paulson@12867
   389
apply (erule disjE)
paulson@12867
   390
 apply (frule mult_le_mono, assumption+)
paulson@12867
   391
apply (simp add: mult_ac)
paulson@12867
   392
apply (auto dest!: natify_eqE 
paulson@12867
   393
            simp add: not_lt_iff_le prime_into_nat mult_le_cancel_le1)
paulson@12867
   394
apply (simp add: prime_def)
paulson@12867
   395
apply (blast dest: lt_trans1)
paulson@12867
   396
done
paulson@12867
   397
paulson@12867
   398
lemma rearrange: "j #* (p#*j) = k#*k \<Longrightarrow> k#*k = p#*(j#*j)"
paulson@12867
   399
by (simp add: mult_ac)
paulson@12867
   400
paulson@12867
   401
lemma prime_not_square:
paulson@12867
   402
     "\<lbrakk>m \<in> nat; p \<in> prime\<rbrakk> \<Longrightarrow> \<forall>k \<in> nat. 0<k \<longrightarrow> m#*m \<noteq> p#*(k#*k)"
paulson@13339
   403
apply (erule complete_induct, clarify)
paulson@13339
   404
apply (frule prime_dvd_other_side, assumption)
paulson@12867
   405
apply assumption
paulson@12867
   406
apply (erule dvdE)
paulson@12867
   407
apply (simp add: mult_assoc mult_cancel1 prime_nonzero prime_into_nat)
paulson@12867
   408
apply (blast dest: rearrange reduction ltD)
paulson@12867
   409
done
paulson@1792
   410
paulson@1792
   411
end