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