src/HOL/ex/Primes.ML
author paulson
Fri, 26 Sep 1997 10:21:14 +0200
changeset 3718 d78cf498a88c
parent 3495 04739732b13e
child 3919 c036caebfc75
permissions -rw-r--r--
Minor tidying to use Clarify_tac, etc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     1
(*  Title:      HOL/ex/Primes.ML
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     2
    ID:         $Id$
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     3
    Author:     Christophe Tabacznyj and Lawrence C Paulson
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     5
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     6
The "divides" relation, the greatest common divisor and Euclid's algorithm
3377
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
     7
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
     8
See H. Davenport, "The Higher Arithmetic".  6th edition.  (CUP, 1992)
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     9
*)
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    10
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    11
eta_contract:=false;
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    12
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    13
open Primes;
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    14
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    15
(************************************************)
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    16
(** Greatest Common Divisor                    **)
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    17
(************************************************)
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    18
3495
04739732b13e New comments on how to deal with unproved termination conditions
paulson
parents: 3457
diff changeset
    19
(*** Euclid's Algorithm ***)
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    20
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    21
3495
04739732b13e New comments on how to deal with unproved termination conditions
paulson
parents: 3457
diff changeset
    22
(** Prove the termination condition and remove it from the recursion equations
04739732b13e New comments on how to deal with unproved termination conditions
paulson
parents: 3457
diff changeset
    23
    and induction rule **)
04739732b13e New comments on how to deal with unproved termination conditions
paulson
parents: 3457
diff changeset
    24
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    25
Tfl.tgoalw thy [] gcd.rules;
3495
04739732b13e New comments on how to deal with unproved termination conditions
paulson
parents: 3457
diff changeset
    26
by (simp_tac (!simpset addsimps [mod_less_divisor, zero_less_eq]) 1);
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    27
val tc = result();
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    28
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    29
val gcd_eq = tc RS hd gcd.rules;
3270
4a3ab8d43451 TFL induction rule is now curried
paulson
parents: 3242
diff changeset
    30
val gcd_induct = tc RS gcd.induct;
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    31
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    32
goal thy "gcd(m,0) = m";
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    33
by (rtac (gcd_eq RS trans) 1);
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    34
by (Simp_tac 1);
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    35
qed "gcd_0";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    36
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    37
goal thy "!!m. 0<n ==> gcd(m,n) = gcd (n, m mod n)";
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    38
by (rtac (gcd_eq RS trans) 1);
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    39
by (asm_simp_tac (!simpset setloop split_tac [expand_if]) 1);
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    40
qed "gcd_less_0";
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    41
Addsimps [gcd_0, gcd_less_0];
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    42
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    43
goal thy "gcd(m,0) dvd m";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    44
by (Simp_tac 1);
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    45
qed "gcd_0_dvd_m";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    46
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    47
goal thy "gcd(m,0) dvd 0";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    48
by (Simp_tac 1);
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    49
qed "gcd_0_dvd_0";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    50
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    51
(*gcd(m,n) divides m and n.  The conjunctions don't seem provable separately*)
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    52
goal thy "(gcd(m,n) dvd m) & (gcd(m,n) dvd n)";
3301
cdcc4d5602b6 Now the recdef induction rule variables are named u, v, ...
paulson
parents: 3270
diff changeset
    53
by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
2102
41a667d2c3fa Replaced excluded_middle_tac by case_tac
paulson
parents: 1804
diff changeset
    54
by (case_tac "n=0" 1);
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    55
by (ALLGOALS 
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    56
    (asm_simp_tac (!simpset addsimps [mod_less_divisor,zero_less_eq])));
3377
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    57
by (blast_tac (!claset addDs [dvd_mod_imp_dvd]) 1);
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    58
qed "gcd_divides_both";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    59
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    60
(*Maximality: for all m,n,f naturals, 
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    61
                if f divides m and f divides n then f divides gcd(m,n)*)
3377
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    62
goal thy "!!k. (f dvd m) --> (f dvd n) --> f dvd gcd(m,n)";
3301
cdcc4d5602b6 Now the recdef induction rule variables are named u, v, ...
paulson
parents: 3270
diff changeset
    63
by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
2102
41a667d2c3fa Replaced excluded_middle_tac by case_tac
paulson
parents: 1804
diff changeset
    64
by (case_tac "n=0" 1);
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    65
by (ALLGOALS 
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    66
    (asm_simp_tac (!simpset addsimps [dvd_mod, mod_less_divisor,
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    67
				      zero_less_eq])));
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    68
qed_spec_mp "gcd_greatest";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    69
3377
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    70
(*Function gcd yields the Greatest Common Divisor*)
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    71
goalw thy [is_gcd_def] "is_gcd (gcd(m,n)) m n";
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    72
by (asm_simp_tac (!simpset addsimps [gcd_greatest, gcd_divides_both]) 1);
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    73
qed "is_gcd";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    74
3377
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    75
(*uniqueness of GCDs*)
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    76
goalw thy [is_gcd_def] "is_gcd m a b & is_gcd n a b --> m=n";
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    77
by (blast_tac (!claset addIs [dvd_anti_sym]) 1);
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 2102
diff changeset
    78
qed "is_gcd_unique";
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    79
3377
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    80
(*Davenport, page 27*)
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    81
goal thy "k * gcd(m,n) = gcd(k*m, k*n)";
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    82
by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    83
by (case_tac "k=0" 1);
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    84
by (case_tac "n=0" 2);
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    85
by (ALLGOALS 
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    86
    (asm_simp_tac (!simpset addsimps [mod_less_divisor, zero_less_eq,
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    87
				      mod_geq, mod_mult_distrib2])));
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    88
qed "gcd_mult_distrib2";
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    89
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    90
(*This theorem leads immediately to a proof of the uniqueness of factorization.
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    91
  If p divides a product of primes then it is one of those primes.*)
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    92
goalw thy [prime_def] "!!p. [| p: prime; p dvd (m*n) |] ==> p dvd m | p dvd n";
3718
d78cf498a88c Minor tidying to use Clarify_tac, etc.
paulson
parents: 3495
diff changeset
    93
by (Clarify_tac 1);
3377
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    94
by (subgoal_tac "m = gcd(m*p, m*n)" 1);
3457
a8ab7c64817c Ran expandshort
paulson
parents: 3377
diff changeset
    95
by (etac ssubst 1);
a8ab7c64817c Ran expandshort
paulson
parents: 3377
diff changeset
    96
by (rtac gcd_greatest 1);
3377
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    97
by (ALLGOALS (asm_simp_tac (!simpset addsimps [gcd_mult_distrib2 RS sym])));
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    98
(*Now deduce  gcd(p,n)=1  to finish the proof*)
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
    99
by (cut_inst_tac [("m","p"),("n","n")] gcd_divides_both 1);
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
   100
by (fast_tac (!claset addSss (!simpset)) 1);
afa1fedef73f New results including the basis for unique factorization
paulson
parents: 3359
diff changeset
   101
qed "prime_dvd_mult";