src/Doc/Tutorial/Rules/TPrimes.thy
author blanchet
Mon, 05 Sep 2016 10:48:06 +0200
changeset 63785 c882ba741244
parent 58860 fee7cfa69c50
child 67406 23307fd33906
permissions -rw-r--r--
added warning
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 *)
55159
608c157d743d Replacing the theory Library/Binomial by Number_Theory/Binomial
paulson <lp15@cam.ac.uk>
parents: 48985
diff changeset
     5
theory TPrimes 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
11234
6902638af59e quantifier instantiation
paulson
parents: 11080
diff changeset
    11
text {*Now in Basic.thy!
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    12
@{thm[display]"dvd_def"}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    13
\rulename{dvd_def}
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
    14
*}
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    15
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    16
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    17
(*** Euclid's Algorithm ***)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    18
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    19
lemma gcd_0 [simp]: "gcd m 0 = m"
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
    20
apply (simp)
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    21
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    22
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    23
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
    24
apply (simp)
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
    25
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    26
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
    27
declare gcd.simps [simp del]
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    28
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    29
(*gcd(m,n) divides m and n.  The conjunctions don't seem provable separately*)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    30
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
    31
apply (induct_tac m n rule: gcd.induct)
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    32
  --{* @{subgoals[display,indent=0,margin=65]} *}
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    33
apply (case_tac "n=0")
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    34
txt{*subgoals after the case tac
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    35
@{subgoals[display,indent=0,margin=65]}
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
    36
*}
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    37
apply (simp_all) 
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    38
  --{* @{subgoals[display,indent=0,margin=65]} *}
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    39
by (blast dest: dvd_mod_imp_dvd)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    40
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    41
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    42
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    43
text {*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    44
@{thm[display] dvd_mod_imp_dvd}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    45
\rulename{dvd_mod_imp_dvd}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    46
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    47
@{thm[display] dvd_trans}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    48
\rulename{dvd_trans}
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    49
*}
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    50
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    51
lemmas gcd_dvd1 [iff] = gcd_dvd_both [THEN conjunct1]
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
    52
lemmas gcd_dvd2 [iff] = gcd_dvd_both [THEN conjunct2]
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    53
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    54
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    55
text {*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    56
\begin{quote}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    57
@{thm[display] gcd_dvd1}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    58
\rulename{gcd_dvd1}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    59
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    60
@{thm[display] gcd_dvd2}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    61
\rulename{gcd_dvd2}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    62
\end{quote}
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
    63
*}
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    64
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    65
(*Maximality: for all m,n,k naturals, 
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    66
                if k divides m and k divides n then k divides gcd(m,n)*)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    67
lemma gcd_greatest [rule_format]:
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    68
      "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
    69
apply (induct_tac m n rule: gcd.induct)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    70
apply (case_tac "n=0")
10853
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    71
txt{*subgoals after the case tac
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    72
@{subgoals[display,indent=0,margin=65]}
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
    73
*}
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    74
apply (simp_all add: dvd_mod)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    75
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    76
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    77
text {*
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    78
@{thm[display] dvd_mod}
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    79
\rulename{dvd_mod}
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    80
*}
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    81
10853
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    82
(*just checking the claim that case_tac "n" works too*)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    83
lemma "k dvd m \<longrightarrow> k dvd n \<longrightarrow> k dvd gcd m n"
10853
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    84
apply (induct_tac m n rule: gcd.induct)
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    85
apply (case_tac "n")
11080
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    86
apply (simp_all add: dvd_mod)
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    87
done
22855d091249 various revisions in response to comments from Tobias
paulson
parents: 10986
diff changeset
    88
10853
2c64c7991f7c case_tac subgoals
paulson
parents: 10846
diff changeset
    89
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    90
theorem gcd_greatest_iff [iff]: 
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
    91
        "(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
    92
by (blast intro!: gcd_greatest intro: dvd_trans)
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    93
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    94
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    95
(**** The material below was omitted from the book ****)
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
    96
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33750
diff changeset
    97
definition is_gcd :: "[nat,nat,nat] \<Rightarrow> bool" where        (*gcd as a relation*)
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    98
    "is_gcd p m n == p dvd m  \<and>  p dvd n  \<and>
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
    99
                     (ALL d. d dvd m \<and> d dvd n \<longrightarrow> d dvd p)"
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   100
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   101
(*Function gcd yields the Greatest Common Divisor*)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   102
lemma is_gcd: "is_gcd (gcd m n) m n"
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
   103
apply (simp add: is_gcd_def gcd_greatest)
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
   104
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   105
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   106
(*uniqueness of GCDs*)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   107
lemma is_gcd_unique: "\<lbrakk> is_gcd m a b; is_gcd n a b \<rbrakk> \<Longrightarrow> m=n"
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
   108
apply (simp add: is_gcd_def)
33750
0a0d6d79d984 anti_sym -> antisym
nipkow
parents: 27657
diff changeset
   109
apply (blast intro: dvd_antisym)
10846
623141a08705 reformatting, and splitting the end of "Primes" to create "Forward"
paulson
parents: 10795
diff changeset
   110
done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   111
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   112
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   113
text {*
33750
0a0d6d79d984 anti_sym -> antisym
nipkow
parents: 27657
diff changeset
   114
@{thm[display] dvd_antisym}
0a0d6d79d984 anti_sym -> antisym
nipkow
parents: 27657
diff changeset
   115
\rulename{dvd_antisym}
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   116
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   117
\begin{isabelle}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   118
proof\ (prove):\ step\ 1\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   119
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   120
goal\ (lemma\ is_gcd_unique):\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   121
\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
   122
\ 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
   123
\ \ \ \ \ \ \ 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
   124
\ \ \ \ \isasymLongrightarrow \ m\ =\ n
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   125
\end{isabelle}
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
   126
*}
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   127
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   128
lemma gcd_assoc: "gcd (gcd k m) n = gcd k (gcd m n)"
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   129
  apply (rule is_gcd_unique)
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   130
  apply (rule is_gcd)
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
   131
  apply (simp add: is_gcd_def)
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 55159
diff changeset
   132
  apply (blast intro: dvd_trans)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   133
  done
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   134
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   135
text{*
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   136
\begin{isabelle}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   137
proof\ (prove):\ step\ 3\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   138
\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   139
goal\ (lemma\ gcd_assoc):\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   140
gcd\ (gcd\ (k,\ m),\ n)\ =\ gcd\ (k,\ gcd\ (m,\ n))\isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   141
\ 1.\ gcd\ (k,\ gcd\ (m,\ n))\ dvd\ k\ \isasymand \isanewline
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   142
\ \ \ \ 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
   143
\end{isabelle}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   144
*}
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   145
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   146
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   147
lemma gcd_dvd_gcd_mult: "gcd m n dvd gcd (k*m) n"
27657
0efc8b68ee4a (adjusted)
haftmann
parents: 25261
diff changeset
   148
  apply (auto intro: dvd_trans [of _ m])
10295
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   149
  done
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   150
33750
0a0d6d79d984 anti_sym -> antisym
nipkow
parents: 27657
diff changeset
   151
(*This is half of the proof (by dvd_antisym) of*)
25261
3dc292be0b54 recdef -> fun
nipkow
parents: 22097
diff changeset
   152
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
   153
  oops
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   154
8eb12693cead the Rules chapter and theories
paulson
parents:
diff changeset
   155
end