src/ZF/ex/Primes.ML
author wenzelm
Mon, 03 Nov 1997 12:24:13 +0100
changeset 4091 771b1f6422a8
parent 2469 b50b8c0eec01
child 4152 451104c223e2
permissions -rw-r--r--
isatool fixclasimp;
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.ML
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     2
    ID:         $Id$
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     3
    Author:     Christophe Tabacznyj and Lawrence C Paulson
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     5
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     6
The "divides" relation, the greatest common divisor and Euclid's algorithm
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     7
*)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     8
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     9
eta_contract:=false;
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    10
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    11
open Primes;
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    12
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    13
(************************************************)
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    14
(** Divides Relation                           **)
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    15
(************************************************)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    16
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    17
goalw thy [dvd_def] "!!m. m dvd n ==> m:nat & n:nat & (EX k:nat. n = m#*k)";
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    18
by (assume_tac 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    19
qed "dvdD";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    20
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    21
bind_thm ("dvd_imp_nat1", dvdD RS conjunct1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    22
bind_thm ("dvd_imp_nat2", dvdD RS conjunct2 RS conjunct1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    23
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    24
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    25
goalw thy [dvd_def] "!!m. m:nat ==> m dvd 0";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    26
by (fast_tac (claset() addIs [nat_0I, mult_0_right RS sym]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    27
qed "dvd_0_right";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    28
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    29
goalw thy [dvd_def] "!!m. 0 dvd m ==> m = 0";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    30
by (fast_tac (claset() addss (simpset())) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    31
qed "dvd_0_left";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    32
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    33
goalw thy [dvd_def] "!!m. m:nat ==> m dvd m";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    34
by (fast_tac (claset() addIs [nat_1I, mult_1_right RS sym]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    35
qed "dvd_refl";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    36
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    37
goalw thy [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    38
by (fast_tac (claset() addIs [mult_assoc, mult_type] ) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    39
qed "dvd_trans";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    40
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    41
goalw thy [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    42
by (fast_tac (claset() addDs [mult_eq_self_implies_10]
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    43
                    addss (simpset() addsimps [mult_assoc, mult_eq_1_iff])) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    44
qed "dvd_anti_sym";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    45
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    46
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    47
(************************************************)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    48
(** Greatest Common Divisor                    **)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    49
(************************************************)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    50
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    51
(* GCD by Euclid's Algorithm *)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    52
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    53
goalw thy [egcd_def] "!!m. m:nat ==> egcd(m,0) = m";
2034
5079fdf938dd Ran expandshort; used stac instead of ssubst
paulson
parents: 1793
diff changeset
    54
by (stac transrec 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2034
diff changeset
    55
by (Asm_simp_tac 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    56
qed "egcd_0";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    57
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    58
goalw thy [egcd_def]
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    59
    "!!m. [| 0<n; m:nat; n:nat |] ==> egcd(m,n) = egcd(n, m mod n)";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    60
by (res_inst_tac [("P", "%z. ?left(z) = ?right")] (transrec RS ssubst) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    61
by (asm_simp_tac (simpset() addsimps [ltD RS mem_imp_not_eq RS not_sym,
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    62
                                     mod_less_divisor RS ltD]) 1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    63
qed "egcd_lt_0";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    64
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    65
goal thy "!!m. m:nat ==> egcd(m,0) dvd m";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    66
by (asm_simp_tac (simpset() addsimps [egcd_0,dvd_refl]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    67
qed "egcd_0_dvd_m";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    68
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    69
goal thy "!!m. m:nat ==> egcd(m,0) dvd 0";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    70
by (asm_simp_tac (simpset() addsimps [egcd_0,dvd_0_right]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    71
qed "egcd_0_dvd_0";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    72
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    73
goalw thy [dvd_def] "!!k. [| k dvd a; k dvd b |] ==> k dvd (a #+ b)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    74
by (fast_tac (claset() addIs [add_mult_distrib_left RS sym, add_type]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    75
qed "dvd_add";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    76
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    77
goalw thy [dvd_def] "!!k. [| k dvd a; q:nat |] ==> k dvd (q #* a)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    78
by (fast_tac (claset() addIs [mult_left_commute, mult_type]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    79
qed "dvd_mult";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    80
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    81
goal thy "!!k. [| k dvd b; k dvd (a mod b); 0 < b; a:nat |] ==> k dvd a";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    82
by (deepen_tac 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    83
    (claset() addIs [mod_div_equality RS subst]
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    84
           addDs [dvdD]
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    85
           addSIs [dvd_add, dvd_mult, mult_type,mod_type,div_type]) 0 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    86
qed "gcd_ind";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    87
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    88
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    89
(* egcd type *)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    90
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    91
goal thy "!!b. b:nat ==> ALL a:nat. egcd(a,b):nat";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    92
by (etac complete_induct 1);
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    93
by (rtac ballI 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    94
by (excluded_middle_tac "x=0" 1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    95
(* case x = 0 *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    96
by (asm_simp_tac (simpset() addsimps [egcd_0]) 2);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    97
(* case x > 0 *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
    98
by (asm_simp_tac (simpset() addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    99
by (eres_inst_tac [("x","a mod x")] ballE 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2034
diff changeset
   100
by (Asm_simp_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   101
by (asm_full_simp_tac (simpset() addsimps [mod_less_divisor RS ltD, 
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   102
                                          nat_into_Ord RS Ord_0_lt]) 1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   103
qed "egcd_type";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   104
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   105
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   106
(* Property 1: egcd(a,b) divides a and b *)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   107
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   108
goal thy "!!b. b:nat ==> ALL a: nat. (egcd(a,b) dvd a) & (egcd(a,b) dvd b)";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   109
by (res_inst_tac [("i","b")] complete_induct 1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   110
by (assume_tac 1);
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
   111
by (rtac ballI 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   112
by (excluded_middle_tac "x=0" 1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   113
(* case x = 0 *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   114
by (asm_simp_tac (simpset() addsimps [egcd_0,dvd_refl,dvd_0_right]) 2);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   115
(* case x > 0 *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   116
by (asm_simp_tac (simpset() addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   117
by (eres_inst_tac [("x","a mod x")] ballE 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2034
diff changeset
   118
by (Asm_simp_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   119
by (asm_full_simp_tac (simpset() addsimps [mod_less_divisor RS ltD, 
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   120
                                          nat_into_Ord RS Ord_0_lt]) 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   121
by (best_tac (claset() addIs [gcd_ind, nat_into_Ord RS Ord_0_lt]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   122
qed "egcd_prop1";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   123
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   124
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   125
(* if f divides a and b then f divides egcd(a,b) *)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   126
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   127
goalw thy [dvd_def] "!!a. [| f dvd a; f dvd b; 0<b |] ==> f dvd (a mod b)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   128
by (safe_tac (claset() addSIs [mult_type, mod_type]));
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   129
ren "m n" 1;
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   130
by (rtac (zero_lt_mult_iff RS iffD1 RS conjE) 1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   131
by (REPEAT_SOME assume_tac);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   132
by (res_inst_tac 
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   133
    [("x", "(((m div n)#*n #+ m mod n) #- ((f#*m) div (f#*n)) #* n)")] 
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   134
    bexI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   135
by (asm_simp_tac (simpset() addsimps [diff_mult_distrib2, div_cancel,
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
   136
                                     mult_mod_distrib, add_mult_distrib_left,
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
   137
                                     diff_add_inverse]) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2034
diff changeset
   138
by (Asm_simp_tac 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   139
qed "dvd_mod";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   140
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   141
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   142
(* Property 2: for all a,b,f naturals, 
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   143
               if f divides a and f divides b then f divides egcd(a,b)*)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   144
goal thy "!!b. [| b:nat; f:nat |] ==>    \
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   145
\              ALL a. (f dvd a) & (f dvd b) --> f dvd egcd(a,b)";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   146
by (etac complete_induct 1);
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
   147
by (rtac allI 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   148
by (excluded_middle_tac "x=0" 1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   149
(* case x = 0 *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   150
by (asm_simp_tac (simpset() addsimps [egcd_0,dvd_refl,dvd_0_right,
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
   151
                                     dvd_imp_nat2]) 2);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   152
(* case x > 0 *)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   153
by (safe_tac (claset()));
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   154
by (asm_simp_tac (simpset() addsimps [egcd_lt_0, nat_into_Ord RS Ord_0_lt,
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
   155
                                     dvd_imp_nat2]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   156
by (eres_inst_tac [("x","a mod x")] ballE 1);
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   157
by (asm_full_simp_tac 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   158
    (simpset() addsimps [mod_less_divisor RS ltD, dvd_imp_nat2, 
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
   159
                        nat_into_Ord RS Ord_0_lt, egcd_lt_0]) 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   160
by (fast_tac (claset() addSIs [dvd_mod, nat_into_Ord RS Ord_0_lt]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   161
qed "egcd_prop2";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   162
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   163
(* GCD PROOF : GCD exists and egcd fits the definition *)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   164
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   165
goalw thy [gcd_def] "!!b. [| a: nat; b:nat |] ==> gcd(egcd(a,b), a, b)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   166
by (asm_simp_tac (simpset() addsimps [egcd_prop1]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   167
by (fast_tac (claset() addIs [egcd_prop2 RS spec RS mp, dvd_imp_nat1]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   168
qed "gcd";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   169
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   170
(* GCD is unique *)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   171
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   172
goalw thy [gcd_def] "!!a. gcd(m,a,b) & gcd(n,a,b) --> m=n";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2469
diff changeset
   173
by (fast_tac (claset() addIs [dvd_anti_sym]) 1);
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   174
qed "gcd_unique";
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
   175