src/HOL/Divides.ML
author paulson
Fri, 30 May 1997 15:15:57 +0200
changeset 3366 2402c6ab1561
child 3427 e7cef2081106
permissions -rw-r--r--
Moving div and mod from Arith to Divides Moving dvd from ex/Primes to Divides
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3366
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Divides.ML
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     2
    ID:         $Id$
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     5
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     6
The division operators div, mod and the divides relation "dvd"
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     7
*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     8
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
     9
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    10
(** Less-then properties **)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    11
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    12
(*In ordinary notation: if 0<n and n<=m then m-n < m *)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    13
goal Arith.thy "!!m. [| 0<n; ~ m<n |] ==> m - n < m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    14
by (subgoal_tac "0<n --> ~ m<n --> m - n < m" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    15
by (Blast_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    16
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    17
by (ALLGOALS(asm_simp_tac(!simpset addsimps [diff_less_Suc])));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    18
qed "diff_less";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    19
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    20
val wf_less_trans = [eq_reflection, wf_pred_nat RS wf_trancl] MRS 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    21
                    def_wfrec RS trans;
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    22
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    23
(*** Remainder ***)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    24
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    25
goal thy "(%m. m mod n) = wfrec (trancl pred_nat) \
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    26
             \                      (%f j. if j<n then j else f (j-n))";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    27
by (simp_tac (!simpset addsimps [mod_def]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    28
qed "mod_eq";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    29
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    30
goal thy "!!m. m<n ==> m mod n = m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    31
by (rtac (mod_eq RS wf_less_trans) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    32
by (Asm_simp_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    33
qed "mod_less";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    34
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    35
goal thy "!!m. [| 0<n;  ~m<n |] ==> m mod n = (m-n) mod n";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    36
by (rtac (mod_eq RS wf_less_trans) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    37
by (asm_simp_tac (!simpset addsimps [diff_less, cut_apply, less_eq]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    38
qed "mod_geq";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    39
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    40
goal thy "m mod 1 = 0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    41
by (induct_tac "m" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    42
by (ALLGOALS (asm_simp_tac (!simpset addsimps [mod_less, mod_geq])));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    43
qed "mod_1";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    44
Addsimps [mod_1];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    45
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    46
goal thy "!!n. 0<n ==> n mod n = 0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    47
by (asm_simp_tac (!simpset addsimps [mod_less, mod_geq]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    48
qed "mod_self";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    49
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    50
goal thy "!!n. 0<n ==> (m+n) mod n = m mod n";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    51
by (subgoal_tac "(n + m) mod n = (n+m-n) mod n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    52
by (stac (mod_geq RS sym) 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    53
by (ALLGOALS (asm_full_simp_tac (!simpset addsimps [add_commute])));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    54
qed "mod_eq_add";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    55
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    56
goal thy "!!k. [| 0<k; 0<n |] ==> (m mod n)*k = (m*k) mod (n*k)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    57
by (res_inst_tac [("n","m")] less_induct 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    58
by (case_tac "na<n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    59
(*case na<n*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    60
by (asm_simp_tac (!simpset addsimps [mod_less]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    61
(*case n<=na*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    62
by (asm_simp_tac (!simpset addsimps [mod_geq, diff_less, zero_less_mult_iff, 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    63
				     diff_mult_distrib]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    64
qed "mod_mult_distrib";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    65
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    66
goal thy "!!k. [| 0<k; 0<n |] ==> k*(m mod n) = (k*m) mod (k*n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    67
by (res_inst_tac [("n","m")] less_induct 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    68
by (case_tac "na<n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    69
(*case na<n*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    70
by (asm_simp_tac (!simpset addsimps [mod_less]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    71
(*case n<=na*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    72
by (asm_simp_tac (!simpset addsimps [mod_geq, diff_less, zero_less_mult_iff, 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    73
				     diff_mult_distrib2]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    74
qed "mod_mult_distrib2";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    75
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    76
goal thy "!!n. 0<n ==> m*n mod n = 0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    77
by (induct_tac "m" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    78
by (asm_simp_tac (!simpset addsimps [mod_less]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    79
by (dres_inst_tac [("m","m*n")] mod_eq_add 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    80
by (asm_full_simp_tac (!simpset addsimps [add_commute]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    81
qed "mod_mult_self_is_0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    82
Addsimps [mod_mult_self_is_0];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    83
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    84
(*** Quotient ***)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    85
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    86
goal thy "(%m. m div n) = wfrec (trancl pred_nat) \
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    87
                        \            (%f j. if j<n then 0 else Suc (f (j-n)))";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    88
by (simp_tac (!simpset addsimps [div_def]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    89
qed "div_eq";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    90
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    91
goal thy "!!m. m<n ==> m div n = 0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    92
by (rtac (div_eq RS wf_less_trans) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    93
by (Asm_simp_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    94
qed "div_less";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    95
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    96
goal thy "!!M. [| 0<n;  ~m<n |] ==> m div n = Suc((m-n) div n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    97
by (rtac (div_eq RS wf_less_trans) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    98
by (asm_simp_tac (!simpset addsimps [diff_less, cut_apply, less_eq]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
    99
qed "div_geq";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   100
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   101
(*Main Result about quotient and remainder.*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   102
goal thy "!!m. 0<n ==> (m div n)*n + m mod n = m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   103
by (res_inst_tac [("n","m")] less_induct 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   104
by (rename_tac "k" 1);    (*Variable name used in line below*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   105
by (case_tac "k<n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   106
by (ALLGOALS (asm_simp_tac(!simpset addsimps ([add_assoc] @
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   107
                       [mod_less, mod_geq, div_less, div_geq,
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   108
                        add_diff_inverse, diff_less]))));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   109
qed "mod_div_equality";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   110
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   111
goal thy "m div 1 = m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   112
by (induct_tac "m" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   113
by (ALLGOALS (asm_simp_tac (!simpset addsimps [div_less, div_geq])));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   114
qed "div_1";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   115
Addsimps [div_1];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   116
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   117
goal thy "!!n. 0<n ==> n div n = 1";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   118
by (asm_simp_tac (!simpset addsimps [div_less, div_geq]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   119
qed "div_self";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   120
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   121
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   122
(*** Further facts about mod (mainly for the mutilated chess board ***)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   123
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   124
goal thy
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   125
    "!!m n. 0<n ==> \
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   126
\           Suc(m) mod n = (if Suc(m mod n) = n then 0 else Suc(m mod n))";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   127
by (res_inst_tac [("n","m")] less_induct 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   128
by (excluded_middle_tac "Suc(na)<n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   129
(* case Suc(na) < n *)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   130
by (forward_tac [lessI RS less_trans] 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   131
by (asm_simp_tac (!simpset addsimps [mod_less, less_not_refl2 RS not_sym]) 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   132
(* case n <= Suc(na) *)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   133
by (asm_full_simp_tac (!simpset addsimps [not_less_iff_le, mod_geq]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   134
by (etac (le_imp_less_or_eq RS disjE) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   135
by (asm_simp_tac (!simpset addsimps [Suc_diff_n]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   136
by (asm_full_simp_tac (!simpset addsimps [not_less_eq RS sym, 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   137
                                          diff_less, mod_geq]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   138
by (asm_simp_tac (!simpset addsimps [mod_less]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   139
qed "mod_Suc";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   140
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   141
goal thy "!!m n. 0<n ==> m mod n < n";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   142
by (res_inst_tac [("n","m")] less_induct 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   143
by (excluded_middle_tac "na<n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   144
(*case na<n*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   145
by (asm_simp_tac (!simpset addsimps [mod_less]) 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   146
(*case n le na*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   147
by (asm_full_simp_tac (!simpset addsimps [mod_geq, diff_less]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   148
qed "mod_less_divisor";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   149
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   150
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   151
(** Evens and Odds **)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   152
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   153
(*With less_zeroE, causes case analysis on b<2*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   154
AddSEs [less_SucE];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   155
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   156
goal thy "!!k b. b<2 ==> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   157
by (subgoal_tac "k mod 2 < 2" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   158
by (asm_simp_tac (!simpset addsimps [mod_less_divisor]) 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   159
by (asm_simp_tac (!simpset setloop split_tac [expand_if]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   160
by (Blast_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   161
qed "mod2_cases";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   162
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   163
goal thy "Suc(Suc(m)) mod 2 = m mod 2";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   164
by (subgoal_tac "m mod 2 < 2" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   165
by (asm_simp_tac (!simpset addsimps [mod_less_divisor]) 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   166
by (Step_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   167
by (ALLGOALS (asm_simp_tac (!simpset addsimps [mod_Suc])));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   168
qed "mod2_Suc_Suc";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   169
Addsimps [mod2_Suc_Suc];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   170
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   171
goal thy "!!m. m mod 2 ~= 0 ==> m mod 2 = 1";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   172
by (subgoal_tac "m mod 2 < 2" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   173
by (asm_simp_tac (!simpset addsimps [mod_less_divisor]) 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   174
by (safe_tac (!claset addSEs [lessE]));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   175
by (ALLGOALS (blast_tac (!claset addIs [sym])));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   176
qed "mod2_neq_0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   177
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   178
goal thy "(m+m) mod 2 = 0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   179
by (induct_tac "m" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   180
by (simp_tac (!simpset addsimps [mod_less]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   181
by (asm_simp_tac (!simpset addsimps [mod2_Suc_Suc, add_Suc_right]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   182
qed "mod2_add_self";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   183
Addsimps [mod2_add_self];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   184
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   185
Delrules [less_SucE];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   186
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   187
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   188
(*** More division laws ***)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   189
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   190
goal thy "!!n. 0<n ==> m*n div n = m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   191
by (cut_inst_tac [("m", "m*n")] mod_div_equality 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   192
ba 1;
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   193
by (asm_full_simp_tac (!simpset addsimps [mod_mult_self_is_0]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   194
qed "div_mult_self_is_m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   195
Addsimps [div_mult_self_is_m];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   196
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   197
(*Cancellation law for division*)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   198
goal thy "!!k. [| 0<n; 0<k |] ==> (k*m) div (k*n) = m div n";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   199
by (res_inst_tac [("n","m")] less_induct 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   200
by (case_tac "na<n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   201
by (asm_simp_tac (!simpset addsimps [div_less, zero_less_mult_iff, 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   202
                                     mult_less_mono2]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   203
by (subgoal_tac "~ k*na < k*n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   204
by (asm_simp_tac
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   205
     (!simpset addsimps [zero_less_mult_iff, div_geq,
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   206
                         diff_mult_distrib2 RS sym, diff_less]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   207
by (asm_full_simp_tac (!simpset addsimps [not_less_iff_le, 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   208
                                          le_refl RS mult_le_mono]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   209
qed "div_cancel";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   210
Addsimps [div_cancel];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   211
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   212
goal thy "!!k. [| 0<n; 0<k |] ==> (k*m) mod (k*n) = k * (m mod n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   213
by (res_inst_tac [("n","m")] less_induct 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   214
by (case_tac "na<n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   215
by (asm_simp_tac (!simpset addsimps [mod_less, zero_less_mult_iff, 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   216
                                     mult_less_mono2]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   217
by (subgoal_tac "~ k*na < k*n" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   218
by (asm_simp_tac
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   219
     (!simpset addsimps [zero_less_mult_iff, mod_geq,
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   220
                         diff_mult_distrib2 RS sym, diff_less]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   221
by (asm_full_simp_tac (!simpset addsimps [not_less_iff_le, 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   222
                                          le_refl RS mult_le_mono]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   223
qed "mult_mod_distrib";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   224
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   225
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   226
(************************************************)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   227
(** Divides Relation                           **)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   228
(************************************************)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   229
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   230
goalw thy [dvd_def] "m dvd 0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   231
by (fast_tac (!claset addIs [mult_0_right RS sym]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   232
qed "dvd_0_right";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   233
Addsimps [dvd_0_right];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   234
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   235
goalw thy [dvd_def] "!!m. 0 dvd m ==> m = 0";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   236
by (fast_tac (!claset addss !simpset) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   237
qed "dvd_0_left";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   238
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   239
goalw thy [dvd_def] "1 dvd k";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   240
by (Simp_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   241
qed "dvd_1_left";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   242
AddIffs [dvd_1_left];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   243
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   244
goalw thy [dvd_def] "m dvd m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   245
by (fast_tac (!claset addIs [mult_1_right RS sym]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   246
qed "dvd_refl";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   247
Addsimps [dvd_refl];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   248
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   249
goalw thy [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   250
by (fast_tac (!claset addIs [mult_assoc] ) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   251
qed "dvd_trans";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   252
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   253
goalw thy [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   254
by (fast_tac (!claset addDs [mult_eq_self_implies_10]
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   255
                     addss (!simpset addsimps [mult_assoc, mult_eq_1_iff])) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   256
qed "dvd_anti_sym";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   257
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   258
goalw thy [dvd_def] "!!k. [| k dvd m; k dvd n |] ==> k dvd (m + n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   259
by (blast_tac (!claset addIs [add_mult_distrib2 RS sym]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   260
qed "dvd_add";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   261
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   262
goalw thy [dvd_def] "!!k. [| k dvd m; k dvd n |] ==> k dvd (m-n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   263
by (blast_tac (!claset addIs [diff_mult_distrib2 RS sym]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   264
qed "dvd_diff";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   265
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   266
goal thy "!!k. [| k dvd (m-n); k dvd n; n<=m |] ==> k dvd m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   267
be (not_less_iff_le RS iffD2 RS add_diff_inverse RS subst) 1;
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   268
by (blast_tac (!claset addIs [dvd_add]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   269
qed "dvd_diffD";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   270
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   271
goalw thy [dvd_def] "!!k. k dvd n ==> k dvd (m*n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   272
by (blast_tac (!claset addIs [mult_left_commute]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   273
qed "dvd_mult";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   274
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   275
goal thy "!!k. k dvd m ==> k dvd (m*n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   276
by (stac mult_commute 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   277
by (etac dvd_mult 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   278
qed "dvd_mult2";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   279
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   280
(* k dvd (m*k) *)
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   281
AddIffs [dvd_refl RS dvd_mult, dvd_refl RS dvd_mult2];
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   282
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   283
goalw thy [dvd_def] "!!m. [| f dvd m; f dvd n; 0<n |] ==> f dvd (m mod n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   284
by (Step_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   285
by (full_simp_tac (!simpset addsimps [zero_less_mult_iff]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   286
by (res_inst_tac 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   287
    [("x", "(((k div ka)*ka + k mod ka) - ((f*k) div (f*ka)) * ka)")] 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   288
    exI 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   289
by (asm_simp_tac (!simpset addsimps [diff_mult_distrib2, 
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   290
                                     mult_mod_distrib, add_mult_distrib2]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   291
qed "dvd_mod";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   292
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   293
goal thy "!!k. [| k dvd (m mod n); k dvd n; n~=0 |] ==> k dvd m";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   294
by (subgoal_tac "k dvd ((m div n)*n + m mod n)" 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   295
by (asm_simp_tac (!simpset addsimps [dvd_add, dvd_mult]) 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   296
by (asm_full_simp_tac (!simpset addsimps [mod_div_equality, zero_less_eq]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   297
qed "dvd_mod_imp_dvd";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   298
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   299
goalw thy [dvd_def]  "!!k m n. [| (k*m) dvd (k*n); 0<k |] ==> m dvd n";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   300
by (etac exE 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   301
by (asm_full_simp_tac (!simpset addsimps mult_ac) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   302
by (Blast_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   303
qed "dvd_mult_cancel";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   304
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   305
goalw thy [dvd_def] "!!i j. [| i dvd m; j dvd n|] ==> (i*j) dvd (m*n)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   306
by (Step_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   307
by (res_inst_tac [("x","k*ka")] exI 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   308
by (asm_simp_tac (!simpset addsimps mult_ac) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   309
qed "mult_dvd_mono";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   310
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   311
goalw thy [dvd_def] "!!c. (i*j) dvd k ==> i dvd k";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   312
by (full_simp_tac (!simpset addsimps [mult_assoc]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   313
by (Blast_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   314
qed "dvd_mult_left";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   315
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   316
goalw thy [dvd_def] "!!n. [| k dvd n; 0 < n |] ==> k <= n";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   317
by (Step_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   318
by (ALLGOALS (full_simp_tac (!simpset addsimps [zero_less_mult_iff])));
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   319
be conjE 1;
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   320
br le_trans 1;
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   321
br (le_refl RS mult_le_mono) 2;
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   322
by (etac Suc_leI 2);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   323
by (Simp_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   324
qed "dvd_imp_le";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   325
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   326
goalw thy [dvd_def] "!!k. 0<k ==> (k dvd n) = (n mod k = 0)";
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   327
by (Step_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   328
by (stac mult_commute 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   329
by (Asm_simp_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   330
by (eres_inst_tac [("t","n")] (mod_div_equality RS subst) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   331
by (asm_simp_tac (!simpset addsimps [mult_commute]) 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   332
by (Blast_tac 1);
2402c6ab1561 Moving div and mod from Arith to Divides
paulson
parents:
diff changeset
   333
qed "dvd_eq_mod_eq_0";