doc-src/TutorialI/Rules/Primes.thy
author wenzelm
Wed, 12 Jan 2011 14:34:11 +0100
changeset 41518 bcacc58902fa
parent 38767 d8da44a8dd25
child 48611 b34ff75c23a7
permissions -rw-r--r--
disabled experimental treatment of replacement text for now, which leads to odd spacing and strange effects on non-poppler viewers;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     1
(* EXTRACT from HOL/ex/Primes.thy*)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     2
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
     3
(*Euclid's algorithm 
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
     4
  This material now appears AFTER that of Forward.thy *)
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 11234
diff changeset
     5
theory Primes imports Main begin
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     6
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
     7
fun gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
     8
  "gcd m n = (if n=0 then m else gcd n (m mod n))"
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
     9
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    10
36745
403585a89772 unified/simplified Pretty.margin_default;
wenzelm
parents: 35416
diff changeset
    11
ML "Pretty.margin_default := 64"
38767
d8da44a8dd25 proper context for various Thy_Output options, via official configuration options in ML and Isar;
wenzelm
parents: 37216
diff changeset
    12
declare [[thy_output_indent = 5]]  (*that is, Doc/TutorialI/settings.ML*)
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    13
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    14
11234
6902638af59e quantifier instantiation
paulson
parents: 11080
diff changeset
    15
text {*Now in Basic.thy!
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    16
@{thm[display]"dvd_def"}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    17
\rulename{dvd_def}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    18
*};
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    19
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    20
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    21
(*** Euclid's Algorithm ***)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    22
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    23
lemma gcd_0 [simp]: "gcd m 0 = m"
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    24
apply (simp);
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    25
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    26
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    27
lemma gcd_non_0 [simp]: "0<n \<Longrightarrow> gcd m n = gcd n (m mod n)"
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    28
apply (simp)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    29
done;
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    30
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    31
declare gcd.simps [simp del];
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    32
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    33
(*gcd(m,n) divides m and n.  The conjunctions don't seem provable separately*)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    34
lemma gcd_dvd_both: "(gcd m n dvd m) \<and> (gcd m n dvd n)"
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    35
apply (induct_tac m n rule: gcd.induct)
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    36
  --{* @{subgoals[display,indent=0,margin=65]} *}
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    37
apply (case_tac "n=0")
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    38
txt{*subgoals after the case tac
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    39
@{subgoals[display,indent=0,margin=65]}
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    40
*};
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    41
apply (simp_all) 
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    42
  --{* @{subgoals[display,indent=0,margin=65]} *}
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    43
by (blast dest: dvd_mod_imp_dvd)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    44
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    45
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    46
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    47
text {*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    48
@{thm[display] dvd_mod_imp_dvd}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    49
\rulename{dvd_mod_imp_dvd}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    50
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    51
@{thm[display] dvd_trans}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    52
\rulename{dvd_trans}
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    53
*}
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    54
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    55
lemmas gcd_dvd1 [iff] = gcd_dvd_both [THEN conjunct1]
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    56
lemmas gcd_dvd2 [iff] = gcd_dvd_both [THEN conjunct2];
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    57
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    58
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    59
text {*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    60
\begin{quote}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    61
@{thm[display] gcd_dvd1}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    62
\rulename{gcd_dvd1}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    63
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    64
@{thm[display] gcd_dvd2}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    65
\rulename{gcd_dvd2}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    66
\end{quote}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    67
*};
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    68
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    69
(*Maximality: for all m,n,k naturals, 
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    70
                if k divides m and k divides n then k divides gcd(m,n)*)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    71
lemma gcd_greatest [rule_format]:
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    72
      "k dvd m \<longrightarrow> k dvd n \<longrightarrow> k dvd gcd m n"
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    73
apply (induct_tac m n rule: gcd.induct)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    74
apply (case_tac "n=0")
10853
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    75
txt{*subgoals after the case tac
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    76
@{subgoals[display,indent=0,margin=65]}
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    77
*};
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    78
apply (simp_all add: dvd_mod)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    79
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    80
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    81
text {*
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    82
@{thm[display] dvd_mod}
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    83
\rulename{dvd_mod}
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    84
*}
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    85
10853
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    86
(*just checking the claim that case_tac "n" works too*)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    87
lemma "k dvd m \<longrightarrow> k dvd n \<longrightarrow> k dvd gcd m n"
10853
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    88
apply (induct_tac m n rule: gcd.induct)
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    89
apply (case_tac "n")
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    90
apply (simp_all add: dvd_mod)
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    91
done
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    92
10853
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    93
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    94
theorem gcd_greatest_iff [iff]: 
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    95
        "(k dvd gcd m n) = (k dvd m \<and> k dvd n)"
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    96
by (blast intro!: gcd_greatest intro: dvd_trans)
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    97
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    98
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    99
(**** The material below was omitted from the book ****)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
   100
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33750
diff changeset
   101
definition is_gcd :: "[nat,nat,nat] \<Rightarrow> bool" where        (*gcd as a relation*)
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   102
    "is_gcd p m n == p dvd m  \<and>  p dvd n  \<and>
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   103
                     (ALL d. d dvd m \<and> d dvd n \<longrightarrow> d dvd p)"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   104
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   105
(*Function gcd yields the Greatest Common Divisor*)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   106
lemma is_gcd: "is_gcd (gcd m n) m n"
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
   107
apply (simp add: is_gcd_def gcd_greatest);
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
   108
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   109
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   110
(*uniqueness of GCDs*)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   111
lemma is_gcd_unique: "\<lbrakk> is_gcd m a b; is_gcd n a b \<rbrakk> \<Longrightarrow> m=n"
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
   112
apply (simp add: is_gcd_def);
33750
0a0d6d79d984 anti_sym -> antisym
nipkow
parents: 27657
diff changeset
   113
apply (blast intro: dvd_antisym)
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
   114
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   115
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   116
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   117
text {*
33750
0a0d6d79d984 anti_sym -> antisym
nipkow
parents: 27657
diff changeset
   118
@{thm[display] dvd_antisym}
0a0d6d79d984 anti_sym -> antisym
nipkow
parents: 27657
diff changeset
   119
\rulename{dvd_antisym}
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   120
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   121
\begin{isabelle}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   122
proof\ (prove):\ step\ 1\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   123
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   124
goal\ (lemma\ is_gcd_unique):\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   125
\isasymlbrakk is_gcd\ m\ a\ b;\ is_gcd\ n\ a\ b\isasymrbrakk \ \isasymLongrightarrow \ m\ =\ n\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   126
\ 1.\ \isasymlbrakk m\ dvd\ a\ \isasymand \ m\ dvd\ b\ \isasymand \ (\isasymforall d.\ d\ dvd\ a\ \isasymand \ d\ dvd\ b\ \isasymlongrightarrow \ d\ dvd\ m);\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   127
\ \ \ \ \ \ \ n\ dvd\ a\ \isasymand \ n\ dvd\ b\ \isasymand \ (\isasymforall d.\ d\ dvd\ a\ \isasymand \ d\ dvd\ b\ \isasymlongrightarrow \ d\ dvd\ n)\isasymrbrakk \isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   128
\ \ \ \ \isasymLongrightarrow \ m\ =\ n
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   129
\end{isabelle}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   130
*};
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   131
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   132
lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)"
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   133
  apply (rule is_gcd_unique)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   134
  apply (rule is_gcd)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   135
  apply (simp add: is_gcd_def);
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   136
  apply (blast intro: dvd_trans);
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   137
  done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   138
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   139
text{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   140
\begin{isabelle}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   141
proof\ (prove):\ step\ 3\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   142
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   143
goal\ (lemma\ gcd_assoc):\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   144
gcd\ (gcd\ (k,\ m),\ n)\ =\ gcd\ (k,\ gcd\ (m,\ n))\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   145
\ 1.\ gcd\ (k,\ gcd\ (m,\ n))\ dvd\ k\ \isasymand \isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   146
\ \ \ \ gcd\ (k,\ gcd\ (m,\ n))\ dvd\ m\ \isasymand \ gcd\ (k,\ gcd\ (m,\ n))\ dvd\ n
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   147
\end{isabelle}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   148
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   149
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   150
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   151
lemma gcd_dvd_gcd_mult: "gcd m n dvd gcd (k*m) n"
27657
0efc8b68ee4a (adjusted)
haftmann
parents: 25261
diff changeset
   152
  apply (auto intro: dvd_trans [of _ m])
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   153
  done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   154
33750
0a0d6d79d984 anti_sym -> antisym
nipkow
parents: 27657
diff changeset
   155
(*This is half of the proof (by dvd_antisym) of*)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   156
lemma gcd_mult_cancel: "gcd k n = 1 \<Longrightarrow> gcd (k*m) n = gcd m n"
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   157
  oops
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   158
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   159
end