src/ZF/Integ/IntDiv.ML
author paulson
Mon, 21 May 2001 14:52:27 +0200
changeset 11321 01cbbf33779b
parent 10635 b115460e5c50
child 11381 4ab3b7b0938f
permissions -rw-r--r--
the rest of integer division
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     1
(*  Title:      HOL/IntDiv.ML
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     2
    ID:         $Id$
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     5
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     6
The division operators div, mod and the divides relation "dvd"
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     7
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     8
Here is the division algorithm in ML:
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
     9
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    10
    fun posDivAlg (a,b) =
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    11
      if a<b then (0,a)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    12
      else let val (q,r) = posDivAlg(a, 2*b)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    13
	       in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    14
	   end;
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    15
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    16
    fun negDivAlg (a,b) =
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    17
      if 0<=a+b then (~1,a+b)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    18
      else let val (q,r) = negDivAlg(a, 2*b)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    19
	       in  if 0<=r-b then (2*q+1, r-b) else (2*q, r)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    20
	   end;
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    21
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    22
    fun negateSnd (q,r:int) = (q,~r);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    23
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    24
    fun divAlg (a,b) = if 0<=a then 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    25
			  if b>0 then posDivAlg (a,b) 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    26
			   else if a=0 then (0,0)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    27
				else negateSnd (negDivAlg (~a,~b))
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    28
		       else 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    29
			  if 0<b then negDivAlg (a,b)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    30
			  else        negateSnd (posDivAlg (~a,~b));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    31
*)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    32
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
    33
Goal "[| #0 $< k; k \\<in> int |] ==> 0 < zmagnitude(k)";
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    34
by (dtac zero_zless_imp_znegative_zminus 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    35
by (dtac zneg_int_of 2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    36
by (auto_tac (claset(), simpset() addsimps [inst "x" "k" zminus_equation]));  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    37
by (subgoal_tac "0 < zmagnitude($# succ(x))" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    38
by (Asm_full_simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    39
by (asm_full_simp_tac (simpset_of Arith.thy addsimps [zmagnitude_int_of]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    40
qed "zero_lt_zmagnitude";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    41
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    42
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    43
(*** Inequality lemmas involving $#succ(m) ***)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    44
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    45
Goal "(w $< z $+ $# succ(m)) <-> (w $< z $+ $#m | intify(w) = z $+ $#m)"; 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    46
by (auto_tac (claset(),
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    47
	      simpset() addsimps [zless_iff_succ_zadd, zadd_assoc, 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    48
                                  int_of_add RS sym]));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    49
by (res_inst_tac [("x","0")] bexI 3);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    50
by (TRYALL (dtac sym));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    51
by (cut_inst_tac [("m","m")] int_succ_int_1 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    52
by (cut_inst_tac [("m","n")] int_succ_int_1 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    53
by (Asm_full_simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    54
by (eres_inst_tac [("n","x")] natE 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    55
by Auto_tac;
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    56
by (res_inst_tac [("x","succ(x)")] bexI 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    57
by Auto_tac;  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    58
qed "zless_add_succ_iff";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    59
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
    60
Goal "z \\<in> int ==> (w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)";
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    61
by (asm_simp_tac (simpset_of Int.thy addsimps
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    62
                  [not_zless_iff_zle RS iff_sym, zless_add_succ_iff]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    63
by (auto_tac (claset() addIs [zle_anti_sym] addEs [zless_asym],
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    64
	      simpset() addsimps [zless_imp_zle, not_zless_iff_zle]));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    65
qed "lemma";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    66
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    67
Goal "(w $+ $# succ(m) $<= z) <-> (w $+ $#m $< z)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    68
by (cut_inst_tac [("z","intify(z)")] lemma 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    69
by Auto_tac;  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    70
qed "zadd_succ_zle_iff";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    71
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    72
(** Inequality reasoning **)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    73
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    74
Goal "(w $< z $+ #1) <-> (w$<=z)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    75
by (subgoal_tac "#1 = $# 1" 1);
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
    76
by (asm_simp_tac (simpset_of Int.thy 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
    77
                  addsimps [zless_add_succ_iff, zle_def]) 1);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    78
by Auto_tac;  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    79
qed "zless_add1_iff_zle";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    80
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    81
Goal "(w $+ #1 $<= z) <-> (w $< z)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    82
by (subgoal_tac "#1 = $# 1" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    83
by (asm_simp_tac (simpset_of Int.thy addsimps [zadd_succ_zle_iff]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    84
by Auto_tac;  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    85
qed "add1_zle_iff";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    86
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    87
Goal "(#1 $+ w $<= z) <-> (w $< z)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    88
by (stac zadd_commute 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    89
by (rtac add1_zle_iff 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    90
qed "add1_left_zle_iff";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    91
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    92
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    93
(*** Monotonicity results ***)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    94
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    95
Goal "(v$+z $< w$+z) <-> (v $< w)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    96
by (Simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    97
qed "zadd_right_cancel_zless";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    98
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
    99
Goal "(z$+v $< z$+w) <-> (v $< w)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   100
by (Simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   101
qed "zadd_left_cancel_zless";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   102
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   103
Addsimps [zadd_right_cancel_zless, zadd_left_cancel_zless];
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   104
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   105
Goal "(v$+z $<= w$+z) <-> (v $<= w)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   106
by (Simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   107
qed "zadd_right_cancel_zle";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   108
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   109
Goal "(z$+v $<= z$+w) <-> (v $<= w)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   110
by (Simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   111
qed "zadd_left_cancel_zle";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   112
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   113
Addsimps [zadd_right_cancel_zle, zadd_left_cancel_zle];
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   114
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   115
(*"v $<= w ==> v$+z $<= w$+z"*)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   116
bind_thm ("zadd_zless_mono1", zadd_right_cancel_zless RS iffD2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   117
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   118
(*"v $<= w ==> z$+v $<= z$+w"*)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   119
bind_thm ("zadd_zless_mono2", zadd_left_cancel_zless RS iffD2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   120
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   121
(*"v $<= w ==> v$+z $<= w$+z"*)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   122
bind_thm ("zadd_zle_mono1", zadd_right_cancel_zle RS iffD2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   123
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   124
(*"v $<= w ==> z$+v $<= z$+w"*)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   125
bind_thm ("zadd_zle_mono2", zadd_left_cancel_zle RS iffD2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   126
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   127
Goal "[| w' $<= w; z' $<= z |] ==> w' $+ z' $<= w $+ z";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   128
by (etac (zadd_zle_mono1 RS zle_trans) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   129
by (Simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   130
qed "zadd_zle_mono";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   131
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   132
Goal "[| w' $< w; z' $<= z |] ==> w' $+ z' $< w $+ z";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   133
by (etac (zadd_zless_mono1 RS zless_zle_trans) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   134
by (Simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   135
qed "zadd_zless_mono";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   136
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   137
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   138
(*** Monotonicity of Multiplication ***)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   139
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   140
Goal "k \\<in> nat ==> i $<= j ==> i $* $#k $<= j $* $#k";
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   141
by (induct_tac "k" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   142
by (stac int_succ_int_1 2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   143
by (ALLGOALS 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   144
    (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2, zadd_zle_mono])));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   145
val lemma = result();
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   146
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   147
Goal "[| i $<= j;  #0 $<= k |] ==> i$*k $<= j$*k";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   148
by (subgoal_tac "i $* intify(k) $<= j $* intify(k)" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   149
by (Full_simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   150
by (res_inst_tac [("b", "intify(k)")] (not_zneg_mag RS subst) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   151
by (rtac lemma 3);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   152
by Auto_tac;
9883
c1c8647af477 a number of new theorems
paulson
parents: 9648
diff changeset
   153
by (asm_full_simp_tac (simpset() addsimps [znegative_iff_zless_0, 
c1c8647af477 a number of new theorems
paulson
parents: 9648
diff changeset
   154
                                           not_zless_iff_zle RS iff_sym]) 1);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   155
qed "zmult_zle_mono1";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   156
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   157
Goal "[| i $<= j;  k $<= #0 |] ==> j$*k $<= i$*k";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   158
by (rtac (zminus_zle_zminus RS iffD1) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   159
by (asm_simp_tac (simpset() delsimps [zmult_zminus_right]
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   160
                            addsimps [zmult_zminus_right RS sym,
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   161
				      zmult_zle_mono1, zle_zminus]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   162
qed "zmult_zle_mono1_neg";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   163
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   164
Goal "[| i $<= j;  #0 $<= k |] ==> k$*i $<= k$*j";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   165
by (dtac zmult_zle_mono1 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   166
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute])));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   167
qed "zmult_zle_mono2";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   168
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   169
Goal "[| i $<= j;  k $<= #0 |] ==> k$*j $<= k$*i";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   170
by (dtac zmult_zle_mono1_neg 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   171
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute])));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   172
qed "zmult_zle_mono2_neg";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   173
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   174
(* $<= monotonicity, BOTH arguments*)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   175
Goal "[| i $<= j;  k $<= l;  #0 $<= j;  #0 $<= k |] ==> i$*k $<= j$*l";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   176
by (etac (zmult_zle_mono1 RS zle_trans) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   177
by (assume_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   178
by (etac zmult_zle_mono2 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   179
by (assume_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   180
qed "zmult_zle_mono";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   181
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   182
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   183
(** strict, in 1st argument; proof is by induction on k>0 **)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   184
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   185
Goal "[| i$<j; k \\<in> nat |] ==> 0<k --> $#k $* i $< $#k $* j";
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   186
by (induct_tac "k" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   187
by (stac int_succ_int_1 2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   188
by (etac natE 2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   189
by (ALLGOALS (asm_full_simp_tac
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   190
	      (simpset() addsimps [zadd_zmult_distrib, zadd_zless_mono, 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   191
				   zle_def])));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   192
by (ftac nat_0_le 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   193
by (mp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   194
by (subgoal_tac "i $+ (i $+ $# xa $* i) $< j $+ (j $+ $# xa $* j)" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   195
by (Full_simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   196
by (rtac zadd_zless_mono 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   197
by (ALLGOALS (asm_simp_tac (simpset() addsimps [zle_def])));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   198
val lemma = result() RS mp;
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   199
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   200
Goal "[| i$<j;  #0 $< k |] ==> k$*i $< k$*j";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   201
by (subgoal_tac "intify(k) $* i $< intify(k) $* j" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   202
by (Full_simp_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   203
by (res_inst_tac [("b", "intify(k)")] (not_zneg_mag RS subst) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   204
by (rtac lemma 3);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   205
by Auto_tac;
9883
c1c8647af477 a number of new theorems
paulson
parents: 9648
diff changeset
   206
by (asm_full_simp_tac (simpset() addsimps [znegative_iff_zless_0]) 1);
c1c8647af477 a number of new theorems
paulson
parents: 9648
diff changeset
   207
by (dtac zless_trans 1 THEN assume_tac 1);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   208
by (auto_tac (claset(), simpset() addsimps [zero_lt_zmagnitude]));  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   209
qed "zmult_zless_mono2";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   210
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   211
Goal "[| i$<j;  #0 $< k |] ==> i$*k $< j$*k";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   212
by (dtac zmult_zless_mono2 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   213
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute])));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   214
qed "zmult_zless_mono1";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   215
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   216
(* < monotonicity, BOTH arguments*)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   217
Goal "[| i $< j;  k $< l;  #0 $< j;  #0 $< k |] ==> i$*k $< j$*l";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   218
by (etac (zmult_zless_mono1 RS zless_trans) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   219
by (assume_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   220
by (etac zmult_zless_mono2 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   221
by (assume_tac 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   222
qed "zmult_zless_mono";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   223
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   224
Goal "[| i $< j;  k $< #0 |] ==> j$*k $< i$*k";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   225
by (rtac (zminus_zless_zminus RS iffD1) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   226
by (asm_simp_tac (simpset() delsimps [zmult_zminus_right]
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   227
                            addsimps [zmult_zminus_right RS sym,
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   228
				      zmult_zless_mono1, zless_zminus]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   229
qed "zmult_zless_mono1_neg";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   230
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   231
Goal "[| i $< j;  k $< #0 |] ==> k$*j $< k$*i";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   232
by (rtac (zminus_zless_zminus RS iffD1) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   233
by (asm_simp_tac (simpset() delsimps [zmult_zminus]
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   234
                            addsimps [zmult_zminus RS sym,
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   235
				      zmult_zless_mono2, zless_zminus]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   236
qed "zmult_zless_mono2_neg";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   237
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   238
Goal "[| m \\<in> int; n \\<in> int |] ==> (m$*n = #0) <-> (m = #0 | n = #0)";
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   239
by (case_tac "m $< #0" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   240
by (auto_tac (claset(), 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   241
     simpset() addsimps [not_zless_iff_zle, zle_def, neq_iff_zless])); 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   242
by (REPEAT 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   243
    (force_tac (claset() addDs [zmult_zless_mono1_neg, zmult_zless_mono1], 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   244
		simpset()) 1));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   245
val lemma = result();
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   246
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   247
Goal "(m$*n = #0) <-> (intify(m) = #0 | intify(n) = #0)";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   248
by (asm_full_simp_tac (simpset() addsimps [lemma RS iff_sym]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   249
qed "zmult_eq_0_iff";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   250
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   251
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   252
(** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =,
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   253
    but not (yet?) for k*m < n*k. **)
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   254
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   255
Goal "[| k \\<in> int; m \\<in> int; n \\<in> int |] \
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   256
\     ==> (m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   257
by (case_tac "k = #0" 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   258
by (auto_tac (claset(), simpset() addsimps [neq_iff_zless, 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   259
                              zmult_zless_mono1, zmult_zless_mono1_neg]));  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   260
by (auto_tac (claset(), 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   261
      simpset() addsimps [not_zless_iff_zle,
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   262
			  inst "w1" "m$*k" (not_zle_iff_zless RS iff_sym),
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   263
			  inst "w1" "m" (not_zle_iff_zless RS iff_sym)]));
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   264
by (ALLGOALS (etac notE));
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   265
by (auto_tac (claset(), simpset() addsimps [zless_imp_zle, zmult_zle_mono1,
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   266
                                            zmult_zle_mono1_neg]));  
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   267
val lemma = result();
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   268
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   269
Goal "(m$*k $< n$*k) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   270
by (cut_inst_tac [("k","intify(k)"),("m","intify(m)"),("n","intify(n)")]
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   271
                 lemma 1);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   272
by Auto_tac;  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   273
qed "zmult_zless_cancel2";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   274
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   275
Goal "(k$*m $< k$*n) <-> ((#0 $< k & m$<n) | (k $< #0 & n$<m))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   276
by (simp_tac (simpset() addsimps [inst "z" "k" zmult_commute, 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   277
                                  zmult_zless_cancel2]) 1);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   278
qed "zmult_zless_cancel1";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   279
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   280
Goal "(m$*k $<= n$*k) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   281
by (simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym, 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   282
                                  zmult_zless_cancel2]) 1);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   283
by Auto_tac;  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   284
qed "zmult_zle_cancel2";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   285
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   286
Goal "(k$*m $<= k$*n) <-> ((#0 $< k --> m$<=n) & (k $< #0 --> n$<=m))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   287
by (simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym, 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   288
                                  zmult_zless_cancel1]) 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   289
by Auto_tac;  
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   290
qed "zmult_zle_cancel1";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   291
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   292
Goal "[| m \\<in> int; n \\<in> int |] ==> m=n <-> (m $<= n & n $<= m)";
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   293
by (blast_tac (claset() addIs [zle_refl,zle_anti_sym]) 1); 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   294
qed "int_eq_iff_zle";
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   295
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   296
Goal "[| k \\<in> int; m \\<in> int; n \\<in> int |] ==> (m$*k = n$*k) <-> (k=#0 | m=n)";
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   297
by (asm_simp_tac (simpset() addsimps [inst "m" "m$*k" int_eq_iff_zle,
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   298
                                      inst "m" "m" int_eq_iff_zle]) 1); 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   299
by (auto_tac (claset(), 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   300
              simpset() addsimps [zmult_zle_cancel2, neq_iff_zless]));  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   301
val lemma = result();
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   302
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   303
Goal "(m$*k = n$*k) <-> (intify(k) = #0 | intify(m) = intify(n))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   304
by (rtac iff_trans 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   305
by (rtac lemma 2);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   306
by Auto_tac;  
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   307
qed "zmult_cancel2";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   308
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   309
Goal "(k$*m = k$*n) <-> (intify(k) = #0 | intify(m) = intify(n))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   310
by (simp_tac (simpset() addsimps [inst "z" "k" zmult_commute, 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   311
                                  zmult_cancel2]) 1);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   312
qed "zmult_cancel1";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   313
Addsimps [zmult_cancel1, zmult_cancel2];
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   314
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   315
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   316
(*** Uniqueness and monotonicity of quotients and remainders ***)
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   317
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   318
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   319
Goal "[| b$*q' $+ r' $<= b$*q $+ r;  #0 $<= r';  #0 $< b;  r $< b |] \
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   320
\     ==> q' $<= q";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   321
by (subgoal_tac "r' $+ b $* (q'$-q) $<= r" 1);
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   322
by (full_simp_tac 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   323
    (simpset() addsimps [zdiff_zmult_distrib2]@zadd_ac@zcompare_rls) 2);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   324
by (subgoal_tac "#0 $< b $* (#1 $+ q $- q')" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   325
by (etac zle_zless_trans 2);
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   326
by (full_simp_tac 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   327
    (simpset() addsimps [zdiff_zmult_distrib2,
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   328
			 zadd_zmult_distrib2]@zadd_ac@zcompare_rls) 2);
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   329
by (etac zle_zless_trans 2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   330
by (Asm_simp_tac 2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   331
by (subgoal_tac "b $* q' $< b $* (#1 $+ q)" 1);
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   332
by (full_simp_tac
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   333
    (simpset() addsimps [zdiff_zmult_distrib2,
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   334
			 zadd_zmult_distrib2]@zadd_ac@zcompare_rls) 2);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   335
by (auto_tac (claset() addEs [zless_asym], 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   336
              simpset() addsimps [zmult_zless_cancel1, zless_add1_iff_zle]@
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9578
diff changeset
   337
                                 zadd_ac@zcompare_rls));
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   338
qed "unique_quotient_lemma";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   339
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   340
Goal "[| b$*q' $+ r' $<= b$*q $+ r;  r $<= #0;  b $< #0;  b $< r' |] \
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   341
\     ==> q $<= q'";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   342
by (res_inst_tac [("b", "$-b"), ("r", "$-r'"), ("r'", "$-r")] 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   343
    unique_quotient_lemma 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   344
by (auto_tac (claset(), 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   345
	      simpset() delsimps [zminus_zadd_distrib]
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   346
			addsimps [zminus_zadd_distrib RS sym,
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   347
	                          zle_zminus, zless_zminus])); 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   348
qed "unique_quotient_lemma_neg";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   349
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   350
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   351
Goal "[| quorem (<a,b>, <q,r>);  quorem (<a,b>, <q',r'>);  b \\<in> int; b ~= #0; \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   352
\        q \\<in> int; q' \\<in> int |] ==> q = q'";
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   353
by (asm_full_simp_tac 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   354
    (simpset() addsimps split_ifs@
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   355
                        [quorem_def, neq_iff_zless]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   356
by Safe_tac; 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   357
by (ALLGOALS Asm_full_simp_tac);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   358
by (REPEAT 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   359
    (blast_tac (claset() addIs [zle_anti_sym]
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   360
			 addDs [zle_eq_refl RS unique_quotient_lemma, 
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   361
				zle_eq_refl RS unique_quotient_lemma_neg,
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   362
				sym]) 1));
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   363
qed "unique_quotient";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   364
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   365
Goal "[| quorem (<a,b>, <q,r>);  quorem (<a,b>, <q',r'>);  b \\<in> int; b ~= #0; \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   366
\        q \\<in> int; q' \\<in> int; \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   367
\        r \\<in> int; r' \\<in> int |] ==> r = r'";
9578
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   368
by (subgoal_tac "q = q'" 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   369
by (blast_tac (claset() addIs [unique_quotient]) 2);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   370
by (asm_full_simp_tac (simpset() addsimps [quorem_def]) 1);
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   371
by Auto_tac;  
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   372
qed "unique_remainder";
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   373
ab26d6c8ebfe new theory Integ/IntDiv and many more monotonicity laws, etc., for the integers
paulson
parents:
diff changeset
   374
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   375
(*** Correctness of posDivAlg, the division algorithm for a>=0 and b>0 ***)
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   376
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   377
Goal "adjust(a, b, <q,r>) = (let diff = r$-b in \
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   378
\                         if #0 $<= diff then <#2$*q $+ #1,diff>  \
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   379
\                                       else <#2$*q,r>)";
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   380
by (simp_tac (simpset() addsimps [Let_def,adjust_def]) 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   381
qed "adjust_eq";
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   382
Addsimps [adjust_eq];
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   383
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   384
10635
b115460e5c50 unsymbolize;
wenzelm
parents: 9955
diff changeset
   385
Goal "[| #0 $< b; \\<not> a $< b |]   \
b115460e5c50 unsymbolize;
wenzelm
parents: 9955
diff changeset
   386
\     ==> nat_of(a $- #2 $\\<times> b $+ #1) < nat_of(a $- b $+ #1)";
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   387
by (simp_tac (simpset() addsimps [zless_nat_conj]) 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   388
by (asm_full_simp_tac (simpset() addsimps [not_zless_iff_zle,
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   389
                       zless_add1_iff_zle]@zcompare_rls) 1); 
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   390
qed "posDivAlg_termination";
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   391
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   392
val posDivAlg_unfold = wf_measure RS (posDivAlg_def RS def_wfrec);
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   393
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   394
Goal "[| #0 $< b; a \\<in> int; b \\<in> int |] ==> \
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   395
\     posDivAlg(<a,b>) =      \
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   396
\      (if a$<b then <#0,a> else adjust(a, b, posDivAlg (<a, #2$*b>)))";
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   397
by (rtac (posDivAlg_unfold RS trans) 1);
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   398
by (asm_simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym]) 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   399
by (asm_simp_tac (simpset() addsimps [vimage_iff, posDivAlg_termination]) 1); 
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   400
qed "posDivAlg_eqn";
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   401
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   402
val [prem] =
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   403
Goal "[| !!a b. [| a \\<in> int; b \\<in> int; \
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   404
\                  ~ (a $< b | b $<= #0) --> P(<a, #2 $* b>) |] \
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   405
\               ==> P(<a,b>) |] \
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   406
\     ==> <u,v> \\<in> int*int --> P(<u,v>)"; 
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   407
by (res_inst_tac [("a","<u,v>")] wf_induct 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   408
by (res_inst_tac [("A","int*int"), ("f","%<a,b>.nat_of (a $- b $+ #1)")] 
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   409
                 wf_measure 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   410
by (Clarify_tac 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   411
by (rtac prem 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   412
by (dres_inst_tac [("x","<xa, #2 $\\<times> y>")] spec 3); 
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   413
by Auto_tac;  
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   414
by (asm_full_simp_tac (simpset() addsimps [not_zle_iff_zless, 
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   415
                                           posDivAlg_termination]) 1); 
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   416
val lemma = result() RS mp;
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   417
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   418
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   419
val prems =
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   420
Goal "[| u \\<in> int; v \\<in> int; \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   421
\        !!a b. [| a \\<in> int; b \\<in> int; ~ (a $< b | b $<= #0) --> P(a, #2 $* b) |] \
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   422
\             ==> P(a,b) |] \
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   423
\     ==> P(u,v)";
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   424
by (subgoal_tac "(%<x,y>. P(x,y))(<u,v>)" 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   425
by (Asm_full_simp_tac 1); 
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   426
by (rtac lemma 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   427
by (simp_tac (simpset() addsimps prems) 2);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   428
by (Full_simp_tac 1);  
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   429
by (resolve_tac prems 1);
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   430
by Auto_tac;  
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   431
qed "posDivAlg_induct";
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
   432
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   433
(*FIXME: use intify in integ_of so that we always have integ_of w \\<in> int.
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   434
    then this rewrite can work for ALL constants!!*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   435
Goal "intify(m) = #0 <-> (m $<= #0 & #0 $<= m)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   436
by (simp_tac (simpset() addsimps [int_eq_iff_zle]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   437
qed "intify_eq_0_iff_zle";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   438
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   439
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   440
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   441
(*** Products of zeroes ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   442
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   443
Goal "[| x \\<in> int; y \\<in> int |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   444
\     ==> (x $* y = #0) <-> (intify(x) = #0 | intify(y) = #0)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   445
by (case_tac "x $< #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   446
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   447
      simpset() addsimps [not_zless_iff_zle, zle_def, neq_iff_zless]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   448
by (REPEAT
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   449
    (force_tac (claset() addDs [zmult_zless_mono1_neg, zmult_zless_mono1], 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   450
		simpset()) 1));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   451
qed "zmult_eq_0_iff_lemma";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   452
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   453
Goal "(x $* y = #0) <-> (intify(x) = #0 | intify(y) = #0)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   454
by (cut_inst_tac [("x","intify(x)"),("y","intify(y)")] 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   455
                 zmult_eq_0_iff_lemma 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   456
by Auto_tac; 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   457
qed "zmult_eq_0_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   458
AddIffs [zmult_eq_0_iff];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   459
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   460
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   461
(*** Some convenient biconditionals for products of signs ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   462
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   463
Goal "[| #0 $< i; #0 $< j |] ==> #0 $< i $* j";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   464
by (dtac zmult_zless_mono1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   465
by Auto_tac; 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   466
qed "zmult_pos";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   467
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   468
Goal "[| i $< #0; j $< #0 |] ==> #0 $< i $* j";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   469
by (dtac zmult_zless_mono1_neg 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   470
by Auto_tac; 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   471
qed "zmult_neg";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   472
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   473
Goal "[| #0 $< i; j $< #0 |] ==> i $* j $< #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   474
by (dtac zmult_zless_mono1_neg 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   475
by Auto_tac; 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   476
qed "zmult_pos_neg";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   477
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   478
(** Inequality reasoning **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   479
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   480
Goal "[| x \\<in> int; y \\<in> int |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   481
\     ==> (#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   482
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   483
              simpset() addsimps [zle_def, not_zless_iff_zle,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   484
	                          zmult_pos, zmult_neg]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   485
by (ALLGOALS (rtac ccontr)); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   486
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   487
	      simpset() addsimps [zle_def, not_zless_iff_zle]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   488
by (ALLGOALS (eres_inst_tac [("P","#0$< x$* y")] rev_mp)); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   489
by (ALLGOALS (dtac zmult_pos_neg THEN' assume_tac));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   490
by (auto_tac (claset() addDs [zless_not_sym], 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   491
              simpset() addsimps [zmult_commute]));  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   492
val lemma = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   493
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   494
Goal "(#0 $< x $* y) <-> (#0 $< x & #0 $< y | x $< #0 & y $< #0)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   495
by (cut_inst_tac [("x","intify(x)"),("y","intify(y)")] lemma 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   496
by Auto_tac; 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   497
qed "int_0_less_mult_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   498
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   499
Goal "[| x \\<in> int; y \\<in> int |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   500
\     ==> (#0 $<= x $* y) <-> (#0 $<= x & #0 $<= y | x $<= #0 & y $<= #0)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   501
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   502
              simpset() addsimps [zle_def, not_zless_iff_zle,  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   503
                                  int_0_less_mult_iff]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   504
val lemma = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   505
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   506
Goal "(#0 $<= x $* y) <-> ((#0 $<= x & #0 $<= y) | (x $<= #0 & y $<= #0))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   507
by (cut_inst_tac [("x","intify(x)"),("y","intify(y)")] lemma 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   508
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   509
qed "int_0_le_mult_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   510
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   511
Goal "(x $* y $< #0) <-> (#0 $< x & y $< #0 | x $< #0 & #0 $< y)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   512
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   513
              simpset() addsimps [int_0_le_mult_iff, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   514
                                  not_zle_iff_zless RS iff_sym]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   515
by (auto_tac (claset() addDs [zless_not_sym],  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   516
              simpset() addsimps [not_zle_iff_zless]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   517
qed "zmult_less_0_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   518
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   519
Goal "(x $* y $<= #0) <-> (#0 $<= x & y $<= #0 | x $<= #0 & #0 $<= y)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   520
by (auto_tac (claset() addDs [zless_not_sym], 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   521
              simpset() addsimps [int_0_less_mult_iff, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   522
                                  not_zless_iff_zle RS iff_sym]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   523
qed "zmult_le_0_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   524
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   525
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   526
(*Typechecking for posDivAlg*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   527
Goal "[| a \\<in> int; b \\<in> int |] ==> posDivAlg(<a,b>) \\<in> int * int";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   528
by (res_inst_tac [("u", "a"), ("v", "b")] posDivAlg_induct 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   529
by (TRYALL assume_tac);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   530
by (case_tac "#0 $< ba" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   531
by (asm_simp_tac (simpset() addsimps [posDivAlg_eqn,adjust_def,integ_of_type]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   532
                            addsplits [split_if_asm]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   533
by (Clarify_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   534
by (asm_full_simp_tac 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   535
    (simpset() addsimps [int_0_less_mult_iff, not_zle_iff_zless]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   536
by (asm_full_simp_tac (simpset() addsimps [not_zless_iff_zle]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   537
by (stac posDivAlg_unfold 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   538
by (Asm_full_simp_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   539
qed_spec_mp "posDivAlg_type";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   540
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   541
(*Correctness of posDivAlg: it computes quotients correctly*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   542
Goal "[| a \\<in> int; b \\<in> int |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   543
\     ==> #0 $<= a --> #0 $< b --> quorem (<a,b>, posDivAlg(<a,b>))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   544
by (res_inst_tac [("u", "a"), ("v", "b")] posDivAlg_induct 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   545
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   546
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [quorem_def])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   547
(*base case: a<b*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   548
by (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff]) 3); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   549
by (asm_full_simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym]) 2); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   550
by (asm_full_simp_tac (simpset() addsimps [posDivAlg_eqn]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   551
(*main argument*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   552
by (stac posDivAlg_eqn 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   553
by (ALLGOALS Asm_simp_tac);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   554
by (etac splitE 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   555
by (rtac posDivAlg_type 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   556
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   557
by (auto_tac (claset(), simpset() addsimps [zadd_zmult_distrib2, Let_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   558
(*now just linear arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   559
by (asm_full_simp_tac 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   560
    (simpset() addsimps [not_zle_iff_zless, zdiff_zless_iff]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   561
qed_spec_mp "posDivAlg_correct";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   562
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   563
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   564
(*** Correctness of negDivAlg, the division algorithm for a<0 and b>0 ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   565
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   566
Goal "[| #0 $< b; \\<not> #0 $<= a $+ b |]   \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   567
\     ==> nat_of($- a $- #2 $\\<times> b) < nat_of($- a $- b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   568
by (simp_tac (simpset() addsimps [zless_nat_conj]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   569
by (asm_full_simp_tac (simpset() addsimps zcompare_rls @ 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   570
           [not_zle_iff_zless, zless_zdiff_iff RS iff_sym, zless_zminus]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   571
qed "negDivAlg_termination";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   572
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   573
val negDivAlg_unfold = wf_measure RS (negDivAlg_def RS def_wfrec);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   574
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   575
Goal "[| #0 $< b; a \\<in> int; b \\<in> int |] ==> \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   576
\     negDivAlg(<a,b>) =      \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   577
\      (if #0 $<= a$+b then <#-1,a$+b> \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   578
\                      else adjust(a, b, negDivAlg (<a, #2$*b>)))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   579
by (rtac (negDivAlg_unfold RS trans) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   580
by (asm_simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   581
by (asm_simp_tac (simpset() addsimps [not_zle_iff_zless, vimage_iff, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   582
                                      negDivAlg_termination]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   583
qed "negDivAlg_eqn";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   584
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   585
val [prem] =
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   586
Goal "[| !!a b. [| a \\<in> int; b \\<in> int; \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   587
\                  ~ (#0 $<= a $+ b | b $<= #0) --> P(<a, #2 $* b>) |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   588
\               ==> P(<a,b>) |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   589
\     ==> <u,v> \\<in> int*int --> P(<u,v>)"; 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   590
by (res_inst_tac [("a","<u,v>")] wf_induct 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   591
by (res_inst_tac [("A","int*int"), ("f","%<a,b>.nat_of ($- a $- b)")] 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   592
                 wf_measure 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   593
by (Clarify_tac 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   594
by (rtac prem 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   595
by (dres_inst_tac [("x","<xa, #2 $\\<times> y>")] spec 3); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   596
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   597
by (asm_full_simp_tac (simpset() addsimps [not_zle_iff_zless, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   598
                                           negDivAlg_termination]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   599
val lemma = result() RS mp;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   600
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   601
val prems =
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   602
Goal "[| u \\<in> int; v \\<in> int; \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   603
\        !!a b. [| a \\<in> int; b \\<in> int; \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   604
\                  ~ (#0 $<= a $+ b | b $<= #0) --> P(a, #2 $* b) |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   605
\               ==> P(a,b) |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   606
\     ==> P(u,v)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   607
by (subgoal_tac "(%<x,y>. P(x,y))(<u,v>)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   608
by (Asm_full_simp_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   609
by (rtac lemma 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   610
by (simp_tac (simpset() addsimps prems) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   611
by (Full_simp_tac 1);  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   612
by (resolve_tac prems 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   613
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   614
qed "negDivAlg_induct";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   615
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   616
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   617
(*Typechecking for negDivAlg*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   618
Goal "[| a \\<in> int; b \\<in> int |] ==> negDivAlg(<a,b>) \\<in> int * int";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   619
by (res_inst_tac [("u", "a"), ("v", "b")] negDivAlg_induct 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   620
by (TRYALL assume_tac);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   621
by (case_tac "#0 $< ba" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   622
by (asm_simp_tac (simpset() addsimps [negDivAlg_eqn,adjust_def,integ_of_type]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   623
                            addsplits [split_if_asm]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   624
by (Clarify_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   625
by (asm_full_simp_tac 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   626
    (simpset() addsimps [int_0_less_mult_iff, not_zle_iff_zless]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   627
by (asm_full_simp_tac (simpset() addsimps [not_zless_iff_zle]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   628
by (stac negDivAlg_unfold 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   629
by (Asm_full_simp_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   630
qed "negDivAlg_type";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   631
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   632
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   633
(*Correctness of negDivAlg: it computes quotients correctly
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   634
  It doesn't work if a=0 because the 0/b=0 rather than -1*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   635
Goal "[| a \\<in> int; b \\<in> int |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   636
\     ==> a $< #0 --> #0 $< b --> quorem (<a,b>, negDivAlg(<a,b>))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   637
by (res_inst_tac [("u", "a"), ("v", "b")] negDivAlg_induct 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   638
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   639
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [quorem_def])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   640
(*base case: 0$<=a$+b*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   641
by (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff]) 3); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   642
by (asm_full_simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym]) 2); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   643
by (asm_full_simp_tac (simpset() addsimps [negDivAlg_eqn]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   644
(*main argument*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   645
by (stac negDivAlg_eqn 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   646
by (ALLGOALS Asm_simp_tac);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   647
by (etac splitE 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   648
by (rtac negDivAlg_type 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   649
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   650
by (auto_tac (claset(), simpset() addsimps [zadd_zmult_distrib2, Let_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   651
(*now just linear arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   652
by (asm_full_simp_tac 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   653
    (simpset() addsimps [not_zle_iff_zless, zdiff_zless_iff]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   654
qed_spec_mp "negDivAlg_correct";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   655
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   656
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   657
(*** Existence shown by proving the division algorithm to be correct ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   658
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   659
(*the case a=0*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   660
Goal "[|b \\<noteq> #0;  b \\<in> int|] ==> quorem (<#0,b>, <#0,#0>)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   661
by (rotate_tac ~1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   662
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   663
	      simpset() addsimps [quorem_def, neq_iff_zless]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   664
qed "quorem_0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   665
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   666
Goal "posDivAlg(<a,#0>) = <#0,a>";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   667
by (stac posDivAlg_unfold 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   668
by (Simp_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   669
qed "posDivAlg_zero_divisor";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   670
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   671
Goal "posDivAlg (<#0,b>) = <#0,#0>";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   672
by (stac posDivAlg_unfold 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   673
by (simp_tac (simpset() addsimps [not_zle_iff_zless]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   674
qed "posDivAlg_0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   675
Addsimps [posDivAlg_0];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   676
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   677
Goal "negDivAlg (<#-1,b>) = <#-1, b$-#1>";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   678
by (stac negDivAlg_unfold 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   679
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   680
(*ALL the rest is linear arithmetic: to notice the contradiction*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   681
by (asm_full_simp_tac (simpset() addsimps [not_zle_iff_zless]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   682
by (dtac (zminus_zless_zminus RS iffD2) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   683
by (asm_full_simp_tac (simpset() addsimps [zadd_commute, zless_add1_iff_zle, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   684
                                           zle_zminus]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   685
by (asm_full_simp_tac (simpset() addsimps [not_zle_iff_zless RS iff_sym]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   686
qed "negDivAlg_minus1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   687
Addsimps [negDivAlg_minus1];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   688
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   689
Goalw [negateSnd_def] "negateSnd (<q,r>) = <q, $-r>";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   690
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   691
qed "negateSnd_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   692
Addsimps [negateSnd_eq];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   693
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   694
Goalw [negateSnd_def] "qr \\<in> int * int ==> negateSnd (qr) \\<in> int * int";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   695
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   696
qed "negateSnd_type";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   697
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   698
Goal "[|quorem (<$-a,$-b>, qr);  a \\<in> int;  b \\<in> int;  qr \\<in> int * int|]  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   699
\     ==> quorem (<a,b>, negateSnd(qr))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   700
by (Clarify_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   701
by (auto_tac (claset() addEs [zless_asym], 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   702
              simpset() addsimps [quorem_def, zless_zminus]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   703
(*linear arithmetic from here on*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   704
by (ALLGOALS
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   705
    (asm_full_simp_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   706
     (simpset() addsimps [inst "x" "a" zminus_equation, zminus_zless])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   707
by (ALLGOALS (cut_inst_tac [("z","b"),("w","#0")] zless_linear));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   708
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   709
by (ALLGOALS (blast_tac (claset() addDs [zle_zless_trans]))); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   710
qed "quorem_neg";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   711
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   712
Goal "[|b \\<noteq> #0;  a \\<in> int;  b \\<in> int|] ==> quorem (<a,b>, divAlg(<a,b>))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   713
by (rotate_tac 1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   714
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   715
	      simpset() addsimps [quorem_0, divAlg_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   716
by (REPEAT_FIRST (ares_tac [quorem_neg, posDivAlg_correct, negDivAlg_correct,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   717
                            posDivAlg_type, negDivAlg_type]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   718
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   719
	      simpset() addsimps [quorem_def, neq_iff_zless]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   720
(*linear arithmetic from here on*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   721
by (auto_tac (claset(), simpset() addsimps [zle_def]));  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   722
qed "divAlg_correct";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   723
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   724
Goal "[|a \\<in> int;  b \\<in> int|] ==> divAlg(<a,b>) \\<in> int * int";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   725
by (auto_tac (claset(), simpset() addsimps [divAlg_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   726
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   727
      simpset() addsimps [posDivAlg_type, negDivAlg_type, negateSnd_type]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   728
qed "divAlg_type";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   729
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   730
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   731
(** intify cancellation **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   732
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   733
Goal "intify(x) zdiv y = x zdiv y";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   734
by (simp_tac (simpset() addsimps [zdiv_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   735
qed "zdiv_intify1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   736
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   737
Goal "x zdiv intify(y) = x zdiv y";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   738
by (simp_tac (simpset() addsimps [zdiv_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   739
qed "zdiv_intify2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   740
Addsimps [zdiv_intify1, zdiv_intify2];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   741
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   742
Goalw [zdiv_def] "z zdiv w \\<in> int";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   743
by (blast_tac (claset() addIs [fst_type, divAlg_type]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   744
qed "zdiv_type";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   745
AddIffs [zdiv_type];  AddTCs [zdiv_type];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   746
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   747
Goal "intify(x) zmod y = x zmod y";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   748
by (simp_tac (simpset() addsimps [zmod_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   749
qed "zmod_intify1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   750
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   751
Goal "x zmod intify(y) = x zmod y";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   752
by (simp_tac (simpset() addsimps [zmod_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   753
qed "zmod_intify2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   754
Addsimps [zmod_intify1, zmod_intify2];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   755
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   756
Goalw [zmod_def] "z zmod w \\<in> int";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   757
by (rtac snd_type 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   758
by (blast_tac (claset() addIs [divAlg_type]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   759
qed "zmod_type";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   760
AddIffs [zmod_type];  AddTCs [zmod_type];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   761
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   762
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   763
(** Arbitrary definitions for division by zero.  Useful to simplify 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   764
    certain equations **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   765
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   766
Goal "a zdiv #0 = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   767
by (simp_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   768
    (simpset() addsimps [zdiv_def, divAlg_def, posDivAlg_zero_divisor]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   769
qed "DIVISION_BY_ZERO_ZDIV";  (*NOT for adding to default simpset*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   770
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   771
Goal "a zmod #0 = intify(a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   772
by (simp_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   773
    (simpset() addsimps [zmod_def, divAlg_def, posDivAlg_zero_divisor]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   774
qed "DIVISION_BY_ZERO_ZMOD";  (*NOT for adding to default simpset*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   775
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   776
fun zdiv_undefined_case_tac s i =
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   777
  case_tac s i THEN 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   778
  asm_simp_tac (simpset() addsimps [DIVISION_BY_ZERO_ZDIV, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   779
				    DIVISION_BY_ZERO_ZMOD]) i;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   780
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   781
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   782
(** Basic laws about division and remainder **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   783
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   784
Goal "[| a \\<in> int; b \\<in> int |] ==> a = b $* (a zdiv b) $+ (a zmod b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   785
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   786
by (cut_inst_tac [("a","a"),("b","b")] divAlg_correct 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   787
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   788
	      simpset() addsimps [quorem_def, zdiv_def, zmod_def, split_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   789
qed "raw_zmod_zdiv_equality";  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   790
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   791
Goal "intify(a) = b $* (a zdiv b) $+ (a zmod b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   792
by (rtac trans 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   793
by (res_inst_tac [("b","intify(b)")] raw_zmod_zdiv_equality 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   794
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   795
qed "zmod_zdiv_equality";  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   796
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   797
Goal "#0 $< b ==> #0 $<= a zmod b & a zmod b $< b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   798
by (cut_inst_tac [("a","intify(a)"),("b","intify(b)")] divAlg_correct 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   799
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   800
	      simpset() addsimps [intify_eq_0_iff_zle, quorem_def, zmod_def, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   801
                                  split_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   802
by (ALLGOALS (blast_tac (claset() addDs [zle_zless_trans]))); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   803
bind_thm ("pos_mod_sign", result() RS conjunct1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   804
bind_thm ("pos_mod_bound", result() RS conjunct2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   805
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   806
Goal "b $< #0 ==> a zmod b $<= #0 & b $< a zmod b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   807
by (cut_inst_tac [("a","intify(a)"),("b","intify(b)")] divAlg_correct 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   808
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   809
	      simpset() addsimps [intify_eq_0_iff_zle, quorem_def, zmod_def, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   810
                                  split_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   811
by (blast_tac (claset() addDs [zle_zless_trans]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   812
by (ALLGOALS (blast_tac (claset() addDs [zless_trans]))); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   813
bind_thm ("neg_mod_sign", result() RS conjunct1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   814
bind_thm ("neg_mod_bound", result() RS conjunct2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   815
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   816
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   817
(** proving general properties of zdiv and zmod **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   818
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   819
Goal "[|b \\<noteq> #0;  a \\<in> int;  b \\<in> int |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   820
\     ==> quorem (<a,b>, <a zdiv b, a zmod b>)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   821
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   822
by (rotate_tac 1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   823
by (auto_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   824
    (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   825
     simpset() addsimps [quorem_def, neq_iff_zless, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   826
			 pos_mod_sign,pos_mod_bound,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   827
			 neg_mod_sign, neg_mod_bound]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   828
qed "quorem_div_mod";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   829
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   830
(*Surely quorem(<a,b>,<q,r>) implies a \\<in> int, but it doesn't matter*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   831
Goal "[| quorem(<a,b>,<q,r>);  b \\<noteq> #0;  a \\<in> int;  b \\<in> int;  q \\<in> int |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   832
\     ==> a zdiv b = q";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   833
by (blast_tac (claset() addIs [quorem_div_mod RS unique_quotient]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   834
qed "quorem_div";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   835
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   836
Goal "[| quorem(<a,b>,<q,r>);  b \\<noteq> #0;  a \\<in> int;  b \\<in> int;  q \\<in> int;  r \\<in> int |] ==> a zmod b = r";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   837
by (blast_tac (claset() addIs [quorem_div_mod RS unique_remainder]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   838
qed "quorem_mod";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   839
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   840
Goal "[| a \\<in> int;  b \\<in> int;  #0 $<= a;  a $< b |] ==> a zdiv b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   841
by (rtac quorem_div 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   842
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   843
(*linear arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   844
by (ALLGOALS (blast_tac (claset() addDs [zle_zless_trans])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   845
qed "zdiv_pos_pos_trivial_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   846
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   847
Goal "[| #0 $<= a;  a $< b |] ==> a zdiv b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   848
by (cut_inst_tac [("a", "intify(a)"), ("b", "intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   849
    zdiv_pos_pos_trivial_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   850
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   851
qed "zdiv_pos_pos_trivial";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   852
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   853
Goal "[| a \\<in> int;  b \\<in> int;  a $<= #0;  b $< a |] ==> a zdiv b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   854
by (res_inst_tac [("r","a")] quorem_div 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   855
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   856
(*linear arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   857
by (ALLGOALS (blast_tac (claset() addDs [zle_zless_trans, zless_trans])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   858
qed "zdiv_neg_neg_trivial_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   859
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   860
Goal "[| a $<= #0;  b $< a |] ==> a zdiv b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   861
by (cut_inst_tac [("a", "intify(a)"), ("b", "intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   862
    zdiv_neg_neg_trivial_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   863
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   864
qed "zdiv_neg_neg_trivial";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   865
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   866
Goal "[| a$+b $<= #0;  #0 $< a;  #0 $< b |] ==> False";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   867
by (dres_inst_tac [("z'","#0"), ("z","b")] zadd_zless_mono 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   868
by (auto_tac (claset(), simpset() addsimps [zle_def]));  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   869
by (blast_tac (claset() addDs [zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   870
qed "zadd_le_0_lemma";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   871
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   872
Goal "[| a \\<in> int;  b \\<in> int;  #0 $< a;  a$+b $<= #0 |] ==> a zdiv b = #-1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   873
by (res_inst_tac [("r","a $+ b ")] quorem_div 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   874
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   875
(*linear arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   876
by (ALLGOALS (blast_tac (claset() addDs [zadd_le_0_lemma, zle_zless_trans])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   877
qed "zdiv_pos_neg_trivial_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   878
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   879
Goal "[| #0 $< a;  a$+b $<= #0 |] ==> a zdiv b = #-1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   880
by (cut_inst_tac [("a", "intify(a)"), ("b", "intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   881
    zdiv_pos_neg_trivial_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   882
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   883
qed "zdiv_pos_neg_trivial";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   884
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   885
(*There is no zdiv_neg_pos_trivial because  #0 zdiv b = #0 would supersede it*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   886
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   887
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   888
Goal "[| a \\<in> int;  b \\<in> int;  #0 $<= a;  a $< b |] ==> a zmod b = a";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   889
by (res_inst_tac [("q","#0")] quorem_mod 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   890
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   891
(*linear arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   892
by (ALLGOALS (blast_tac (claset() addDs [zle_zless_trans])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   893
qed "zmod_pos_pos_trivial_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   894
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   895
Goal "[| #0 $<= a;  a $< b |] ==> a zmod b = intify(a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   896
by (cut_inst_tac [("a", "intify(a)"), ("b", "intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   897
    zmod_pos_pos_trivial_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   898
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   899
qed "zmod_pos_pos_trivial";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   900
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   901
Goal "[| a \\<in> int;  b \\<in> int;  a $<= #0;  b $< a |] ==> a zmod b = a";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   902
by (res_inst_tac [("q","#0")] quorem_mod 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   903
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   904
(*linear arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   905
by (ALLGOALS (blast_tac (claset() addDs [zle_zless_trans, zless_trans])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   906
qed "zmod_neg_neg_trivial_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   907
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   908
Goal "[| a $<= #0;  b $< a |] ==> a zmod b = intify(a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   909
by (cut_inst_tac [("a", "intify(a)"), ("b", "intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   910
    zmod_neg_neg_trivial_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   911
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   912
qed "zmod_neg_neg_trivial";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   913
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   914
Goal "[| a \\<in> int;  b \\<in> int;  #0 $< a;  a$+b $<= #0 |] ==> a zmod b = a$+b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   915
by (res_inst_tac [("q","#-1")] quorem_mod 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   916
by (auto_tac (claset(), simpset() addsimps [quorem_def]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   917
(*linear arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   918
by (ALLGOALS (blast_tac (claset() addDs [zadd_le_0_lemma, zle_zless_trans])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   919
qed "zmod_pos_neg_trivial_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   920
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   921
Goal "[| #0 $< a;  a$+b $<= #0 |] ==> a zmod b = a$+b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   922
by (cut_inst_tac [("a", "intify(a)"), ("b", "intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   923
    zmod_pos_neg_trivial_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   924
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   925
qed "zmod_pos_neg_trivial";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   926
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   927
(*There is no zmod_neg_pos_trivial...*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   928
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   929
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   930
(*Simpler laws such as -a zdiv b = -(a zdiv b) FAIL*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   931
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   932
Goal "[|a \\<in> int;  b \\<in> int|] ==> ($-a) zdiv ($-b) = a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   933
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   934
by (stac ((simplify(simpset()) (quorem_div_mod RS quorem_neg)) 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   935
	  RS quorem_div) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   936
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   937
qed "zdiv_zminus_zminus_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   938
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   939
Goal "($-a) zdiv ($-b) = a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   940
by (cut_inst_tac [("a","intify(a)"), ("b","intify(b)")] 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   941
    zdiv_zminus_zminus_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   942
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   943
qed "zdiv_zminus_zminus";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   944
Addsimps [zdiv_zminus_zminus];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   945
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   946
(*Simpler laws such as -a zmod b = -(a zmod b) FAIL*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   947
Goal "[|a \\<in> int;  b \\<in> int|] ==> ($-a) zmod ($-b) = $- (a zmod b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   948
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   949
by (stac ((simplify(simpset()) (quorem_div_mod RS quorem_neg)) 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   950
	  RS quorem_mod) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   951
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   952
qed "zmod_zminus_zminus_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   953
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   954
Goal "($-a) zmod ($-b) = $- (a zmod b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   955
by (cut_inst_tac [("a","intify(a)"), ("b","intify(b)")] 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   956
    zmod_zminus_zminus_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   957
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   958
qed "zmod_zminus_zminus";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   959
Addsimps [zmod_zminus_zminus];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   960
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   961
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   962
(*** division of a number by itself ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   963
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   964
Goal "[| #0 $< a; a = r $+ a$*q; r $< a |] ==> #1 $<= q";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   965
by (subgoal_tac "#0 $< a$*q" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   966
by (cut_inst_tac [("w","#0"),("z","q")] add1_zle_iff 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   967
by (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   968
by (blast_tac (claset() addDs [zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   969
(*linear arithmetic...*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   970
by (dres_inst_tac [("t","%x. x $- r")] subst_context 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   971
by (dtac sym 1);  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   972
by (asm_full_simp_tac (simpset() addsimps zcompare_rls) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   973
val lemma1 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   974
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   975
Goal "[| #0 $< a; a = r $+ a$*q; #0 $<= r |] ==> q $<= #1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   976
by (subgoal_tac "#0 $<= a$*(#1$-q)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   977
by (asm_simp_tac (simpset() addsimps [zdiff_zmult_distrib2]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   978
by (dres_inst_tac [("t","%x. x $- a $* q")] subst_context 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   979
by (asm_full_simp_tac (simpset() addsimps zcompare_rls) 2); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   980
by (asm_full_simp_tac (simpset() addsimps int_0_le_mult_iff::zcompare_rls) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   981
by (blast_tac (claset() addDs [zle_zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   982
val lemma2 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   983
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   984
Goal "[| quorem(<a,a>,<q,r>);  a \\<in> int;  q \\<in> int;  a \\<noteq> #0|] ==> q = #1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   985
by (asm_full_simp_tac 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   986
    (simpset() addsimps split_ifs@[quorem_def, neq_iff_zless]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   987
by (rtac zle_anti_sym 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   988
by Safe_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   989
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   990
by (blast_tac (claset() addDs [zless_trans]) 4); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   991
by (blast_tac (claset() addDs [zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   992
by (res_inst_tac [("a", "$-a"),("r", "$-r")] lemma1 3);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   993
by (res_inst_tac [("a", "$-a"),("r", "$-r")] lemma2 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   994
by (rtac (zminus_equation RS iffD1) 6); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   995
by (rtac (zminus_equation RS iffD1) 2); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   996
by (REPEAT (force_tac  (claset() addIs [lemma1,lemma2], 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   997
	      simpset() addsimps [zadd_commute, zmult_zminus]) 1));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   998
qed "self_quotient";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
   999
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1000
Goal "[|quorem(<a,a>,<q,r>); a \\<in> int; q \\<in> int; r \\<in> int; a \\<noteq> #0|] ==> r = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1001
by (ftac self_quotient 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1002
by (auto_tac (claset(), simpset() addsimps [quorem_def]));  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1003
qed "self_remainder";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1004
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1005
Goal "[|a \\<noteq> #0; a \\<in> int|] ==> a zdiv a = #1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1006
by (blast_tac (claset() addIs [quorem_div_mod RS self_quotient]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1007
qed "zdiv_self_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1008
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1009
Goal "intify(a) \\<noteq> #0 ==> a zdiv a = #1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1010
by (dtac zdiv_self_raw 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1011
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1012
qed "zdiv_self";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1013
Addsimps [zdiv_self];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1014
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1015
(*Here we have 0 zmod 0 = 0, also assumed by Knuth (who puts m zmod 0 = 0) *)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1016
Goal "a \\<in> int ==> a zmod a = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1017
by (zdiv_undefined_case_tac "a = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1018
by (blast_tac (claset() addIs [quorem_div_mod RS self_remainder]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1019
qed "zmod_self_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1020
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1021
Goal "a zmod a = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1022
by (cut_inst_tac [("a","intify(a)")] zmod_self_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1023
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1024
qed "zmod_self";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1025
Addsimps [zmod_self];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1026
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1027
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1028
(*** Computation of division and remainder ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1029
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1030
Goal "#0 zdiv b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1031
by (simp_tac (simpset() addsimps [zdiv_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1032
qed "zdiv_zero";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1033
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1034
Goal "#0 $< b ==> #-1 zdiv b = #-1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1035
by (asm_simp_tac (simpset() addsimps [zdiv_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1036
qed "zdiv_eq_minus1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1037
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1038
Goal "#0 zmod b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1039
by (simp_tac (simpset() addsimps [zmod_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1040
qed "zmod_zero";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1041
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1042
Addsimps [zdiv_zero, zmod_zero];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1043
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1044
Goal "#0 $< b ==> #-1 zdiv b = #-1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1045
by (asm_simp_tac (simpset() addsimps [zdiv_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1046
qed "zdiv_minus1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1047
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1048
Goal "#0 $< b ==> #-1 zmod b = b $- #1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1049
by (asm_simp_tac (simpset() addsimps [zmod_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1050
qed "zmod_minus1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1051
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1052
(** a positive, b positive **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1053
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1054
Goal "[| #0 $< a;  #0 $<= b |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1055
\     ==> a zdiv b = fst (posDivAlg(<intify(a), intify(b)>))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1056
by (asm_simp_tac (simpset() addsimps [zdiv_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1057
by (auto_tac (claset(), simpset() addsimps [zle_def]));  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1058
qed "zdiv_pos_pos";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1059
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1060
Goal "[| #0 $< a;  #0 $<= b |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1061
\     ==> a zmod b = snd (posDivAlg(<intify(a), intify(b)>))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1062
by (asm_simp_tac (simpset() addsimps [zmod_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1063
by (auto_tac (claset(), simpset() addsimps [zle_def]));  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1064
qed "zmod_pos_pos";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1065
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1066
(** a negative, b positive **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1067
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1068
Goal "[| a $< #0;  #0 $< b |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1069
\     ==> a zdiv b = fst (negDivAlg(<intify(a), intify(b)>))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1070
by (asm_simp_tac (simpset() addsimps [zdiv_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1071
by (blast_tac (claset() addDs [zle_zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1072
qed "zdiv_neg_pos";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1073
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1074
Goal "[| a $< #0;  #0 $< b |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1075
\     ==> a zmod b = snd (negDivAlg(<intify(a), intify(b)>))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1076
by (asm_simp_tac (simpset() addsimps [zmod_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1077
by (blast_tac (claset() addDs [zle_zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1078
qed "zmod_neg_pos";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1079
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1080
(** a positive, b negative **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1081
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1082
Goal "[| #0 $< a;  b $< #0 |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1083
\     ==> a zdiv b = fst (negateSnd(negDivAlg (<$-a, $-b>)))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1084
by (asm_simp_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1085
    (simpset() addsimps [zdiv_def, divAlg_def, intify_eq_0_iff_zle]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1086
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1087
by (REPEAT (blast_tac (claset() addDs [zle_zless_trans]) 1));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1088
by (blast_tac (claset() addDs [zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1089
by (blast_tac (claset() addIs [zless_imp_zle]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1090
qed "zdiv_pos_neg";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1091
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1092
Goal "[| #0 $< a;  b $< #0 |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1093
\     ==> a zmod b = snd (negateSnd(negDivAlg (<$-a, $-b>)))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1094
by (asm_simp_tac 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1095
    (simpset() addsimps [zmod_def, divAlg_def, intify_eq_0_iff_zle]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1096
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1097
by (REPEAT (blast_tac (claset() addDs [zle_zless_trans]) 1));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1098
by (blast_tac (claset() addDs [zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1099
by (blast_tac (claset() addIs [zless_imp_zle]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1100
qed "zmod_pos_neg";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1101
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1102
(** a negative, b negative **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1103
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1104
Goal "[| a $< #0;  b $<= #0 |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1105
\     ==> a zdiv b = fst (negateSnd(posDivAlg(<$-a, $-b>)))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1106
by (asm_simp_tac (simpset() addsimps [zdiv_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1107
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1108
by (REPEAT (blast_tac (claset() addSDs [zle_zless_trans]) 1));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1109
qed "zdiv_neg_neg";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1110
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1111
Goal "[| a $< #0;  b $<= #0 |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1112
\     ==> a zmod b = snd (negateSnd(posDivAlg(<$-a, $-b>)))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1113
by (asm_simp_tac (simpset() addsimps [zmod_def, divAlg_def]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1114
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1115
by (REPEAT (blast_tac (claset() addSDs [zle_zless_trans]) 1));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1116
qed "zmod_neg_neg";
9955
6ed42bcba707 a bit more of division
paulson
parents: 9883
diff changeset
  1117
11321
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1118
Addsimps (map (read_instantiate_sg (sign_of (the_context ()))
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1119
	       [("a", "integ_of (?v)"), ("b", "integ_of (?w)")])
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1120
	  [zdiv_pos_pos, zdiv_neg_pos, zdiv_pos_neg, zdiv_neg_neg,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1121
	   zmod_pos_pos, zmod_neg_pos, zmod_pos_neg, zmod_neg_neg,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1122
	   posDivAlg_eqn, negDivAlg_eqn]);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1123
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1124
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1125
(** Special-case simplification **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1126
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1127
Goal "a zmod #1 = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1128
by (cut_inst_tac [("a","a"),("b","#1")] pos_mod_sign 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1129
by (cut_inst_tac [("a","a"),("b","#1")] pos_mod_bound 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1130
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1131
(*arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1132
by (dtac (add1_zle_iff RS iffD2) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1133
by (rtac zle_anti_sym 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1134
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1135
qed "zmod_1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1136
Addsimps [zmod_1];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1137
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1138
Goal "a zdiv #1 = intify(a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1139
by (cut_inst_tac [("a","a"),("b","#1")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1140
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1141
qed "zdiv_1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1142
Addsimps [zdiv_1];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1143
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1144
Goal "a zmod #-1 = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1145
by (cut_inst_tac [("a","a"),("b","#-1")] neg_mod_sign 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1146
by (cut_inst_tac [("a","a"),("b","#-1")] neg_mod_bound 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1147
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1148
(*arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1149
by (dtac (add1_zle_iff RS iffD2) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1150
by (rtac zle_anti_sym 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1151
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1152
qed "zmod_minus1_right";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1153
Addsimps [zmod_minus1_right];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1154
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1155
Goal "a \\<in> int ==> a zdiv #-1 = $-a";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1156
by (cut_inst_tac [("a","a"),("b","#-1")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1157
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1158
by (rtac (equation_zminus RS iffD2) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1159
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1160
qed "zdiv_minus1_right_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1161
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1162
Goal "a zdiv #-1 = $-a";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1163
by (cut_inst_tac [("a","intify(a)")] zdiv_minus1_right_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1164
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1165
qed "zdiv_minus1_right";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1166
Addsimps [zdiv_minus1_right];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1167
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1168
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1169
(*** Monotonicity in the first argument (divisor) ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1170
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1171
Goal "[| a $<= a';  #0 $< b |] ==> a zdiv b $<= a' zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1172
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1173
by (cut_inst_tac [("a","a'"),("b","b")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1174
by (rtac unique_quotient_lemma 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1175
by (etac subst 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1176
by (etac subst 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1177
by (ALLGOALS (asm_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1178
qed "zdiv_mono1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1179
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1180
Goal "[| a $<= a';  b $< #0 |] ==> a' zdiv b $<= a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1181
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1182
by (cut_inst_tac [("a","a'"),("b","b")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1183
by (rtac unique_quotient_lemma_neg 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1184
by (etac subst 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1185
by (etac subst 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1186
by (ALLGOALS (asm_simp_tac (simpset() addsimps [neg_mod_sign,neg_mod_bound])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1187
qed "zdiv_mono1_neg";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1188
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1189
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1190
(*** Monotonicity in the second argument (dividend) ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1191
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1192
Goal "[| b$*q $+ r = b'$*q' $+ r';  #0 $<= b'$*q' $+ r';  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1193
\        r' $< b';  #0 $<= r;  #0 $< b';  b' $<= b |]  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1194
\     ==> q $<= q'";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1195
by (subgoal_tac "#0 $<= q'" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1196
 by (subgoal_tac "#0 $< b'$*(q' $+ #1)" 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1197
  by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 3);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1198
  by (etac zle_zless_trans 3); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1199
  by (etac zadd_zless_mono2 3);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1200
 by (asm_full_simp_tac (simpset() addsimps [int_0_less_mult_iff]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1201
 by (blast_tac (claset() addDs [zless_trans]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1202
                         addIs  [zless_add1_iff_zle RS iffD1]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1203
by (subgoal_tac "b$*q $< b$*(q' $+ #1)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1204
 by (asm_full_simp_tac (simpset() addsimps [zmult_zless_cancel1]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1205
 by (force_tac (claset() addDs  [zless_add1_iff_zle RS iffD1,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1206
                                 zless_trans, zless_zle_trans], 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1207
                simpset()) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1208
by (subgoal_tac "b$*q = r' $- r $+ b'$*q'" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1209
 by (asm_simp_tac (simpset() addsimps zcompare_rls) 2); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1210
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1211
by (stac zadd_commute 1 THEN rtac zadd_zless_mono 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1212
 by (blast_tac (claset() addIs [zmult_zle_mono1]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1213
by (subgoal_tac "r' $+ #0 $< b $+ r" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1214
 by (asm_full_simp_tac (simpset() addsimps zcompare_rls) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1215
by (rtac zadd_zless_mono 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1216
 by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1217
by (blast_tac (claset() addDs [zless_zle_trans]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1218
qed "zdiv_mono2_lemma";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1219
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1220
Goal "[| #0 $<= a;  #0 $< b';  b' $<= b;  a \\<in> int |]  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1221
\     ==> a zdiv b $<= a zdiv b'";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1222
by (subgoal_tac "#0 $< b" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1223
 by (blast_tac (claset() addDs [zless_zle_trans]) 2); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1224
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1225
by (cut_inst_tac [("a","a"),("b","b'")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1226
by (rtac zdiv_mono2_lemma 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1227
by (etac subst 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1228
by (etac subst 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1229
by (ALLGOALS
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1230
    (asm_full_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1231
qed "zdiv_mono2_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1232
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1233
Goal "[| #0 $<= a;  #0 $< b';  b' $<= b |]  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1234
\     ==> a zdiv b $<= a zdiv b'";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1235
by (cut_inst_tac [("a","intify(a)")] zdiv_mono2_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1236
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1237
qed "zdiv_mono2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1238
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1239
Goal "[| b$*q $+ r = b'$*q' $+ r';  b'$*q' $+ r' $< #0;  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1240
\        r $< b;  #0 $<= r';  #0 $< b';  b' $<= b |]  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1241
\     ==> q' $<= q";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1242
by (subgoal_tac "#0 $< b" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1243
 by (blast_tac (claset() addDs [zless_zle_trans]) 2); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1244
by (subgoal_tac "q' $< #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1245
 by (subgoal_tac "b'$*q' $< #0" 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1246
  by (force_tac (claset() addIs [zle_zless_trans], simpset()) 3); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1247
 by (asm_full_simp_tac (simpset() addsimps [zmult_less_0_iff]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1248
 by (blast_tac (claset() addDs [zless_trans]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1249
by (subgoal_tac "b$*q' $< b$*(q $+ #1)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1250
 by (asm_full_simp_tac (simpset() addsimps [zmult_zless_cancel1]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1251
 by (blast_tac (claset() addDs [zless_trans, zless_add1_iff_zle RS iffD1]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1252
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1253
by (subgoal_tac "b$*q' $<= b'$*q'" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1254
 by (asm_simp_tac (simpset() addsimps [zmult_zle_cancel2]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1255
 by (blast_tac (claset() addDs [zless_trans]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1256
by (subgoal_tac "b'$*q' $+ r $< b $+ (b$*q $+ r)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1257
 by (etac ssubst 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1258
 by (Asm_simp_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1259
 by (dres_inst_tac [("w'","r"),("z'","#0")] zadd_zless_mono 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1260
  by (assume_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1261
 by (Asm_full_simp_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1262
by (full_simp_tac (simpset() addsimps [zadd_commute]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1263
by (rtac zle_zless_trans 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1264
by (assume_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1265
 by (asm_simp_tac (simpset() addsimps [zmult_zle_cancel2]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1266
by (blast_tac (claset() addDs [zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1267
qed "zdiv_mono2_neg_lemma";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1268
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1269
Goal "[| a $< #0;  #0 $< b';  b' $<= b;  a \\<in> int |]  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1270
\     ==> a zdiv b' $<= a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1271
by (subgoal_tac "#0 $< b" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1272
 by (blast_tac (claset() addDs [zless_zle_trans]) 2); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1273
by (cut_inst_tac [("a","a"),("b","b")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1274
by (cut_inst_tac [("a","a"),("b","b'")] zmod_zdiv_equality 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1275
by (rtac zdiv_mono2_neg_lemma 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1276
by (etac subst 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1277
by (etac subst 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1278
by (ALLGOALS
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1279
    (asm_full_simp_tac (simpset() addsimps [pos_mod_sign,pos_mod_bound])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1280
qed "zdiv_mono2_neg_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1281
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1282
Goal "[| a $< #0;  #0 $< b';  b' $<= b |]  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1283
\     ==> a zdiv b' $<= a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1284
by (cut_inst_tac [("a","intify(a)")] zdiv_mono2_neg_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1285
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1286
qed "zdiv_mono2_neg";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1287
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1288
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1289
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1290
(*** More algebraic laws for zdiv and zmod ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1291
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1292
(** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1293
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1294
Goal "[| quorem(<b,c>, <q,r>);  c \\<in> int;  c \\<noteq> #0 |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1295
\     ==> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1296
by (auto_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1297
    (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1298
     simpset() addsimps split_ifs@
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1299
			[quorem_def, neq_iff_zless, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1300
			 zadd_zmult_distrib2,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1301
			 pos_mod_sign,pos_mod_bound,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1302
			 neg_mod_sign, neg_mod_bound]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1303
by (ALLGOALS (rtac raw_zmod_zdiv_equality));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1304
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1305
qed "zmult1_lemma";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1306
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1307
Goal "[|b \\<in> int;  c \\<in> int|] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1308
\     ==> (a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1309
by (zdiv_undefined_case_tac "c = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1310
by (rtac (quorem_div_mod RS zmult1_lemma RS quorem_div) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1311
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1312
qed "zdiv_zmult1_eq_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1313
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1314
Goal "(a$*b) zdiv c = a$*(b zdiv c) $+ a$*(b zmod c) zdiv c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1315
by (cut_inst_tac [("b","intify(b)"), ("c","intify(c)")] zdiv_zmult1_eq_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1316
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1317
qed "zdiv_zmult1_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1318
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1319
Goal "[|b \\<in> int;  c \\<in> int|] ==> (a$*b) zmod c = a$*(b zmod c) zmod c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1320
by (zdiv_undefined_case_tac "c = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1321
by (rtac (quorem_div_mod RS zmult1_lemma RS quorem_mod) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1322
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1323
qed "zmod_zmult1_eq_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1324
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1325
Goal "(a$*b) zmod c = a$*(b zmod c) zmod c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1326
by (cut_inst_tac [("b","intify(b)"), ("c","intify(c)")] zmod_zmult1_eq_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1327
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1328
qed "zmod_zmult1_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1329
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1330
Goal "(a$*b) zmod c = ((a zmod c) $* b) zmod c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1331
by (rtac trans 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1332
by (res_inst_tac [("b", "(b $* a) zmod c")] trans 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1333
by (rtac zmod_zmult1_eq 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1334
by (ALLGOALS (simp_tac (simpset() addsimps [zmult_commute])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1335
qed "zmod_zmult1_eq'";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1336
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1337
Goal "(a$*b) zmod c = ((a zmod c) $* (b zmod c)) zmod c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1338
by (rtac (zmod_zmult1_eq' RS trans) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1339
by (rtac zmod_zmult1_eq 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1340
qed "zmod_zmult_distrib";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1341
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1342
Goal "intify(b) \\<noteq> #0 ==> (a$*b) zdiv b = intify(a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1343
by (asm_simp_tac (simpset() addsimps [zdiv_zmult1_eq]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1344
qed "zdiv_zmult_self1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1345
Addsimps [zdiv_zmult_self1];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1346
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1347
Goal "intify(b) \\<noteq> #0 ==> (b$*a) zdiv b = intify(a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1348
by (stac zmult_commute 1 THEN etac zdiv_zmult_self1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1349
qed "zdiv_zmult_self2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1350
Addsimps [zdiv_zmult_self2];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1351
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1352
Goal "(a$*b) zmod b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1353
by (simp_tac (simpset() addsimps [zmod_zmult1_eq]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1354
qed "zmod_zmult_self1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1355
Addsimps [zmod_zmult_self1];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1356
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1357
Goal "(b$*a) zmod b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1358
by (simp_tac (simpset() addsimps [zmult_commute, zmod_zmult1_eq]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1359
qed "zmod_zmult_self2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1360
Addsimps [zmod_zmult_self2];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1361
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1362
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1363
(** proving (a$+b) zdiv c = 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1364
            a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1365
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1366
Goal "[| quorem(<a,c>, <aq,ar>);  quorem(<b,c>, <bq,br>);  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1367
\        c \\<in> int;  c \\<noteq> #0 |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1368
\     ==> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1369
by (auto_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1370
    (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1371
     simpset() addsimps split_ifs@
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1372
			[quorem_def, neq_iff_zless, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1373
			 zadd_zmult_distrib2,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1374
			 pos_mod_sign,pos_mod_bound,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1375
			 neg_mod_sign, neg_mod_bound]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1376
by (ALLGOALS (rtac raw_zmod_zdiv_equality));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1377
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1378
val zadd1_lemma = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1379
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1380
(*NOT suitable for rewriting: the RHS has an instance of the LHS*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1381
Goal "[|a \\<in> int; b \\<in> int; c \\<in> int|] ==> \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1382
\     (a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1383
by (zdiv_undefined_case_tac "c = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1384
by (blast_tac (claset() addIs [[quorem_div_mod,quorem_div_mod]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1385
			       MRS zadd1_lemma RS quorem_div]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1386
qed "zdiv_zadd1_eq_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1387
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1388
Goal "(a$+b) zdiv c = a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1389
by (cut_inst_tac [("a","intify(a)"), ("b","intify(b)"), ("c","intify(c)")] 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1390
    zdiv_zadd1_eq_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1391
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1392
qed "zdiv_zadd1_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1393
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1394
Goal "[|a \\<in> int; b \\<in> int; c \\<in> int|]  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1395
\     ==> (a$+b) zmod c = (a zmod c $+ b zmod c) zmod c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1396
by (zdiv_undefined_case_tac "c = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1397
by (blast_tac (claset() addIs [[quorem_div_mod,quorem_div_mod]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1398
			       MRS zadd1_lemma RS quorem_mod]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1399
qed "zmod_zadd1_eq_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1400
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1401
Goal "(a$+b) zmod c = (a zmod c $+ b zmod c) zmod c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1402
by (cut_inst_tac [("a","intify(a)"), ("b","intify(b)"), ("c","intify(c)")] 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1403
    zmod_zadd1_eq_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1404
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1405
qed "zmod_zadd1_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1406
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1407
Goal "[|a \\<in> int; b \\<in> int|] ==> (a zmod b) zdiv b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1408
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1409
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1410
      simpset() addsimps [neq_iff_zless, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1411
			  pos_mod_sign, pos_mod_bound, zdiv_pos_pos_trivial, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1412
			  neg_mod_sign, neg_mod_bound, zdiv_neg_neg_trivial]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1413
qed "zmod_div_trivial_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1414
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1415
Goal "(a zmod b) zdiv b = #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1416
by (cut_inst_tac [("a","intify(a)"), ("b","intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1417
    zmod_div_trivial_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1418
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1419
qed "zmod_div_trivial";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1420
Addsimps [zmod_div_trivial];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1421
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1422
Goal "[|a \\<in> int; b \\<in> int|] ==> (a zmod b) zmod b = a zmod b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1423
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1424
by (auto_tac (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1425
       simpset() addsimps [neq_iff_zless, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1426
			   pos_mod_sign, pos_mod_bound, zmod_pos_pos_trivial, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1427
			   neg_mod_sign, neg_mod_bound, zmod_neg_neg_trivial]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1428
qed "zmod_mod_trivial_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1429
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1430
Goal "(a zmod b) zmod b = a zmod b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1431
by (cut_inst_tac [("a","intify(a)"), ("b","intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1432
    zmod_mod_trivial_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1433
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1434
qed "zmod_mod_trivial";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1435
Addsimps [zmod_mod_trivial];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1436
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1437
Goal "(a$+b) zmod c = ((a zmod c) $+ b) zmod c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1438
by (rtac (trans RS sym) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1439
by (rtac zmod_zadd1_eq 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1440
by (Simp_tac 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1441
by (rtac (zmod_zadd1_eq RS sym) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1442
qed "zmod_zadd_left_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1443
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1444
Goal "(a$+b) zmod c = (a $+ (b zmod c)) zmod c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1445
by (rtac (trans RS sym) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1446
by (rtac zmod_zadd1_eq 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1447
by (Simp_tac 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1448
by (rtac (zmod_zadd1_eq RS sym) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1449
qed "zmod_zadd_right_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1450
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1451
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1452
Goal "intify(a) \\<noteq> #0 ==> (a$+b) zdiv a = b zdiv a $+ #1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1453
by (asm_simp_tac (simpset() addsimps [zdiv_zadd1_eq]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1454
qed "zdiv_zadd_self1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1455
Addsimps [zdiv_zadd_self1];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1456
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1457
Goal "intify(a) \\<noteq> #0 ==> (b$+a) zdiv a = b zdiv a $+ #1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1458
by (asm_simp_tac (simpset() addsimps [zdiv_zadd1_eq]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1459
qed "zdiv_zadd_self2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1460
Addsimps [zdiv_zadd_self2];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1461
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1462
Goal "(a$+b) zmod a = b zmod a";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1463
by (zdiv_undefined_case_tac "a = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1464
by (asm_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1465
qed "zmod_zadd_self1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1466
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1467
Goal "(b$+a) zmod a = b zmod a";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1468
by (zdiv_undefined_case_tac "a = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1469
by (asm_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1470
qed "zmod_zadd_self2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1471
Addsimps [zmod_zadd_self1, zmod_zadd_self2];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1472
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1473
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1474
(*** proving  a zdiv (b*c) = (a zdiv b) zdiv c ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1475
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1476
(*The condition c>0 seems necessary.  Consider that 7 zdiv ~6 = ~2 but
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1477
  7 zdiv 2 zdiv ~3 = 3 zdiv ~3 = ~1.  The subcase (a zdiv b) zmod c = 0 seems
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1478
  to cause particular problems.*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1479
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1480
(** first, four lemmas to bound the remainder for the cases b<0 and b>0 **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1481
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1482
Goal "[| #0 $< c;  b $< r;  r $<= #0 |] ==> b$*c $< b$*(q zmod c) $+ r";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1483
by (subgoal_tac "b $* (c $- q zmod c) $< r $* #1" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1484
by (asm_full_simp_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1485
    (simpset() addsimps [zdiff_zmult_distrib2, zadd_commute]@zcompare_rls) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1486
by (rtac zle_zless_trans 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1487
by (etac zmult_zless_mono1 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1488
by (rtac zmult_zle_mono2_neg 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1489
by (auto_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1490
    (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1491
     simpset() addsimps zcompare_rls@
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1492
			[zadd_commute, add1_zle_iff, pos_mod_bound]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1493
by (blast_tac (claset() addIs [zless_imp_zle]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1494
			addDs [zless_zle_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1495
val lemma1 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1496
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1497
Goal "[| #0 $< c;   b $< r;  r $<= #0 |] ==> b $* (q zmod c) $+ r $<= #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1498
by (subgoal_tac "b $* (q zmod c) $<= #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1499
by (asm_simp_tac (simpset() addsimps [zmult_le_0_iff, pos_mod_sign]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1500
by (blast_tac (claset() addIs [zless_imp_zle]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1501
			addDs [zless_zle_trans]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1502
(*arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1503
by (dtac zadd_zle_mono 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1504
by (assume_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1505
by (asm_full_simp_tac (simpset() addsimps [zadd_commute]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1506
val lemma2 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1507
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1508
Goal "[| #0 $< c;  #0 $<= r;  r $< b |] ==> #0 $<= b $* (q zmod c) $+ r";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1509
by (subgoal_tac "#0 $<= b $* (q zmod c)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1510
by (asm_simp_tac (simpset() addsimps [int_0_le_mult_iff, pos_mod_sign]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1511
by (blast_tac (claset() addIs [zless_imp_zle]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1512
			addDs [zle_zless_trans]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1513
(*arithmetic*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1514
by (dtac zadd_zle_mono 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1515
by (assume_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1516
by (asm_full_simp_tac (simpset() addsimps [zadd_commute]) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1517
val lemma3 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1518
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1519
Goal "[| #0 $< c; #0 $<= r; r $< b |] ==> b $* (q zmod c) $+ r $< b $* c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1520
by (subgoal_tac "r $* #1 $< b $* (c $- q zmod c)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1521
by (asm_full_simp_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1522
    (simpset() addsimps [zdiff_zmult_distrib2, zadd_commute]@zcompare_rls) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1523
by (rtac zless_zle_trans 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1524
by (etac zmult_zless_mono1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1525
by (rtac zmult_zle_mono2 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1526
by (auto_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1527
    (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1528
     simpset() addsimps zcompare_rls@
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1529
			[zadd_commute, add1_zle_iff, pos_mod_bound]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1530
by (blast_tac (claset() addIs [zless_imp_zle]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1531
			addDs [zle_zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1532
val lemma4 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1533
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1534
Goal "[| quorem (<a,b>, <q,r>);  a \\<in> int;  b \\<in> int;  b \\<noteq> #0;  #0 $< c |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1535
\     ==> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1536
by (auto_tac  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1537
    (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1538
     simpset() addsimps zmult_ac@
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1539
			[zmod_zdiv_equality RS sym, quorem_def, neq_iff_zless,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1540
			 int_0_less_mult_iff,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1541
			 zadd_zmult_distrib2 RS sym,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1542
			 lemma1, lemma2, lemma3, lemma4]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1543
by (ALLGOALS (blast_tac (claset() addDs [zless_trans])));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1544
val lemma = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1545
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1546
Goal "[|#0 $< c;  a \\<in> int;  b \\<in> int|] ==> a zdiv (b$*c) = (a zdiv b) zdiv c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1547
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1548
by (rtac (quorem_div_mod RS lemma RS quorem_div) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1549
by (auto_tac (claset(), simpset() addsimps [intify_eq_0_iff_zle]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1550
by (blast_tac (claset() addDs [zle_zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1551
qed "zdiv_zmult2_eq_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1552
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1553
Goal "#0 $< c ==> a zdiv (b$*c) = (a zdiv b) zdiv c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1554
by (cut_inst_tac [("a","intify(a)"), ("b","intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1555
    zdiv_zmult2_eq_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1556
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1557
qed "zdiv_zmult2_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1558
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1559
Goal "[|#0 $< c;  a \\<in> int;  b \\<in> int|] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1560
\     ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1561
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1562
by (rtac (quorem_div_mod RS lemma RS quorem_mod) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1563
by (auto_tac (claset(), simpset() addsimps [intify_eq_0_iff_zle]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1564
by (blast_tac (claset() addDs [zle_zless_trans]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1565
qed "zmod_zmult2_eq_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1566
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1567
Goal "#0 $< c ==> a zmod (b$*c) = b$*(a zdiv b zmod c) $+ a zmod b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1568
by (cut_inst_tac [("a","intify(a)"), ("b","intify(b)")]
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1569
    zmod_zmult2_eq_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1570
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1571
qed "zmod_zmult2_eq";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1572
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1573
(*** Cancellation of common factors in "zdiv" ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1574
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1575
Goal "[| #0 $< b;  intify(c) \\<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1576
by (stac zdiv_zmult2_eq 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1577
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1578
val lemma1 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1579
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1580
Goal "[| b $< #0;  intify(c) \\<noteq> #0 |] ==> (c$*a) zdiv (c$*b) = a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1581
by (subgoal_tac "(c $* ($-a)) zdiv (c $* ($-b)) = ($-a) zdiv ($-b)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1582
by (rtac lemma1 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1583
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1584
val lemma2 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1585
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1586
Goal "[|intify(c) \\<noteq> #0; b \\<in> int|] ==> (c$*a) zdiv (c$*b) = a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1587
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1588
by (auto_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1589
    (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1590
     simpset() addsimps [read_instantiate [("x", "b")] neq_iff_zless, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1591
			 lemma1, lemma2]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1592
qed "zdiv_zmult_zmult1_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1593
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1594
Goal "intify(c) \\<noteq> #0 ==> (c$*a) zdiv (c$*b) = a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1595
by (cut_inst_tac [("b","intify(b)")] zdiv_zmult_zmult1_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1596
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1597
qed "zdiv_zmult_zmult1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1598
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1599
Goal "intify(c) \\<noteq> #0 ==> (a$*c) zdiv (b$*c) = a zdiv b";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1600
by (dtac zdiv_zmult_zmult1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1601
by (auto_tac (claset(), simpset() addsimps [zmult_commute]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1602
qed "zdiv_zmult_zmult2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1603
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1604
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1605
(*** Distribution of factors over "zmod" ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1606
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1607
Goal "[| #0 $< b;  intify(c) \\<noteq> #0 |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1608
\     ==> (c$*a) zmod (c$*b) = c $* (a zmod b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1609
by (stac zmod_zmult2_eq 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1610
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1611
val lemma1 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1612
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1613
Goal "[| b $< #0;  intify(c) \\<noteq> #0 |] \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1614
\     ==> (c$*a) zmod (c$*b) = c $* (a zmod b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1615
by (subgoal_tac "(c $* ($-a)) zmod (c $* ($-b)) = c $* (($-a) zmod ($-b))" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1616
by (rtac lemma1 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1617
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1618
val lemma2 = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1619
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1620
Goal "[|b \\<in> int; c \\<in> int|] ==> (c$*a) zmod (c$*b) = c $* (a zmod b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1621
by (zdiv_undefined_case_tac "b = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1622
by (zdiv_undefined_case_tac "c = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1623
by (auto_tac
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1624
    (claset(), 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1625
     simpset() addsimps [read_instantiate [("x", "b")] neq_iff_zless, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1626
			 lemma1, lemma2]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1627
qed "zmod_zmult_zmult1_raw";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1628
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1629
Goal "(c$*a) zmod (c$*b) = c $* (a zmod b)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1630
by (cut_inst_tac [("b","intify(b)"),("c","intify(c)")] 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1631
    zmod_zmult_zmult1_raw 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1632
by Auto_tac;  
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1633
qed "zmod_zmult_zmult1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1634
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1635
Goal "(a$*c) zmod (b$*c) = (a zmod b) $* c";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1636
by (cut_inst_tac [("c","c")] zmod_zmult_zmult1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1637
by (auto_tac (claset(), simpset() addsimps [zmult_commute]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1638
qed "zmod_zmult_zmult2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1639
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1640
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1641
(** Quotients of signs **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1642
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1643
Goal "[| a $< #0;  #0 $< b |] ==> a zdiv b $< #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1644
by (subgoal_tac "a zdiv b $<= #-1" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1645
by (etac zle_zless_trans 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1646
by (Simp_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1647
by (rtac zle_trans 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1648
by (res_inst_tac [("a'","#-1")]  zdiv_mono1 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1649
by (rtac (zless_add1_iff_zle RS iffD1) 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1650
by (Simp_tac 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1651
by (auto_tac (claset(), simpset() addsimps [zdiv_minus1]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1652
qed "zdiv_neg_pos_less0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1653
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1654
Goal "[| #0 $<= a;  b $< #0 |] ==> a zdiv b $<= #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1655
by (dtac zdiv_mono1_neg 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1656
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1657
qed "zdiv_nonneg_neg_le0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1658
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1659
Goal "#0 $< b ==> (#0 $<= a zdiv b) <-> (#0 $<= a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1660
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1661
by (dtac zdiv_mono1 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1662
by (auto_tac (claset(), simpset() addsimps [neq_iff_zless]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1663
by (full_simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1664
by (blast_tac (claset() addIs [zdiv_neg_pos_less0]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1665
qed "pos_imp_zdiv_nonneg_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1666
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1667
Goal "b $< #0 ==> (#0 $<= a zdiv b) <-> (a $<= #0)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1668
by (stac (zdiv_zminus_zminus RS sym) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1669
by (rtac iff_trans 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1670
by (rtac pos_imp_zdiv_nonneg_iff 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1671
by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1672
qed "neg_imp_zdiv_nonneg_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1673
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1674
(*But not (a zdiv b $<= 0 iff a$<=0); consider a=1, b=2 when a zdiv b = 0.*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1675
Goal "#0 $< b ==> (a zdiv b $< #0) <-> (a $< #0)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1676
by (asm_simp_tac (simpset() addsimps [not_zle_iff_zless RS iff_sym]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1677
by (etac pos_imp_zdiv_nonneg_iff 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1678
qed "pos_imp_zdiv_neg_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1679
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1680
(*Again the law fails for $<=: consider a = -1, b = -2 when a zdiv b = 0*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1681
Goal "b $< #0 ==> (a zdiv b $< #0) <-> (#0 $< a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1682
by (asm_simp_tac (simpset() addsimps [not_zle_iff_zless RS iff_sym]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1683
by (etac neg_imp_zdiv_nonneg_iff 1); 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1684
qed "neg_imp_zdiv_neg_iff";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1685
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1686
(*
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1687
 THESE REMAIN TO BE CONVERTED -- but aren't that useful!
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1688
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1689
 (*** Speeding up the division algorithm with shifting ***)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1690
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1691
 (** computing "zdiv" by shifting **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1692
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1693
 Goal "#0 $<= a ==> (#1 $+ #2$*b) zdiv (#2$*a) = b zdiv a";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1694
 by (zdiv_undefined_case_tac "a = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1695
 by (subgoal_tac "#1 $<= a" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1696
  by (arith_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1697
 by (subgoal_tac "#1 $< a $* #2" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1698
  by (arith_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1699
 by (subgoal_tac "#2$*(#1 $+ b zmod a) $<= #2$*a" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1700
  by (rtac zmult_zle_mono2 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1701
 by (auto_tac (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1702
	       simpset() addsimps [zadd_commute, zmult_commute, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1703
				   add1_zle_iff, pos_mod_bound]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1704
 by (stac zdiv_zadd1_eq 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1705
 by (asm_simp_tac (simpset() addsimps [zdiv_zmult_zmult2, zmod_zmult_zmult2, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1706
				       zdiv_pos_pos_trivial]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1707
 by (stac zdiv_pos_pos_trivial 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1708
 by (asm_simp_tac (simpset() 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1709
	    addsimps [zmod_pos_pos_trivial,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1710
		     pos_mod_sign RS zadd_zle_mono1 RSN (2,zle_trans)]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1711
 by (auto_tac (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1712
	       simpset() addsimps [zmod_pos_pos_trivial]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1713
 by (subgoal_tac "#0 $<= b zmod a" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1714
  by (asm_simp_tac (simpset() addsimps [pos_mod_sign]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1715
 by (arith_tac 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1716
 qed "pos_zdiv_mult_2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1717
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1718
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1719
 Goal "a $<= #0 ==> (#1 $+ #2$*b) zdiv (#2$*a) <-> (b$+#1) zdiv a";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1720
 by (subgoal_tac "(#1 $+ #2$*($-b-#1)) zdiv (#2 $* ($-a)) <-> ($-b-#1) zdiv ($-a)" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1721
 by (rtac pos_zdiv_mult_2 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1722
 by (auto_tac (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1723
	       simpset() addsimps [zmult_zminus_right]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1724
 by (subgoal_tac "(#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1725
 by (Simp_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1726
 by (asm_full_simp_tac (HOL_ss
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1727
			addsimps [zdiv_zminus_zminus, zdiff_def,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1728
				  zminus_zadd_distrib RS sym]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1729
 qed "neg_zdiv_mult_2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1730
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1731
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1732
 (*Not clear why this must be proved separately; probably integ_of causes
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1733
   simplification problems*)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1734
 Goal "~ #0 $<= x ==> x $<= #0";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1735
 by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1736
 val lemma = result();
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1737
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1738
 Goal "integ_of (v BIT b) zdiv integ_of (w BIT False) = \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1739
 \         (if ~b | #0 $<= integ_of w                   \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1740
 \          then integ_of v zdiv (integ_of w)    \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1741
 \          else (integ_of v $+ #1) zdiv (integ_of w))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1742
 by (simp_tac (simpset_of Int.thy addsimps [zadd_assoc, integ_of_BIT]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1743
 by (asm_simp_tac (simpset()
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1744
		   delsimps bin_arith_extra_simps@bin_rel_simps
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1745
		   addsimps [zdiv_zmult_zmult1,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1746
			     pos_zdiv_mult_2, lemma, neg_zdiv_mult_2]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1747
 qed "zdiv_integ_of_BIT";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1748
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1749
 Addsimps [zdiv_integ_of_BIT];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1750
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1751
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1752
 (** computing "zmod" by shifting (proofs resemble those for "zdiv") **)
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1753
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1754
 Goal "#0 $<= a ==> (#1 $+ #2$*b) zmod (#2$*a) = #1 $+ #2 $* (b zmod a)";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1755
 by (zdiv_undefined_case_tac "a = #0" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1756
 by (subgoal_tac "#1 $<= a" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1757
  by (arith_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1758
 by (subgoal_tac "#1 $< a $* #2" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1759
  by (arith_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1760
 by (subgoal_tac "#2$*(#1 $+ b zmod a) $<= #2$*a" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1761
  by (rtac zmult_zle_mono2 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1762
 by (auto_tac (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1763
	       simpset() addsimps [zadd_commute, zmult_commute, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1764
				   add1_zle_iff, pos_mod_bound]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1765
 by (stac zmod_zadd1_eq 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1766
 by (asm_simp_tac (simpset() addsimps [zmod_zmult_zmult2, 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1767
				       zmod_pos_pos_trivial]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1768
 by (rtac zmod_pos_pos_trivial 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1769
 by (asm_simp_tac (simpset() 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1770
 #		  addsimps [zmod_pos_pos_trivial,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1771
		     pos_mod_sign RS zadd_zle_mono1 RSN (2,zle_trans)]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1772
 by (auto_tac (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1773
	       simpset() addsimps [zmod_pos_pos_trivial]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1774
 by (subgoal_tac "#0 $<= b zmod a" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1775
  by (asm_simp_tac (simpset() addsimps [pos_mod_sign]) 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1776
 by (arith_tac 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1777
 qed "pos_zmod_mult_2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1778
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1779
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1780
 Goal "a $<= #0 ==> (#1 $+ #2$*b) zmod (#2$*a) = #2 $* ((b$+#1) zmod a) - #1";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1781
 by (subgoal_tac 
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1782
     "(#1 $+ #2$*($-b-#1)) zmod (#2$*($-a)) = #1 $+ #2$*(($-b-#1) zmod ($-a))" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1783
 by (rtac pos_zmod_mult_2 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1784
 by (auto_tac (claset(),
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1785
	       simpset() addsimps [zmult_zminus_right]));
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1786
 by (subgoal_tac "(#-1 - (#2 $* b)) = - (#1 $+ (#2 $* b))" 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1787
 by (Simp_tac 2);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1788
 by (asm_full_simp_tac (HOL_ss
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1789
			addsimps [zmod_zminus_zminus, zdiff_def,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1790
				  zminus_zadd_distrib RS sym]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1791
 by (dtac (zminus_equation RS iffD1 RS sym) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1792
 by Auto_tac;
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1793
 qed "neg_zmod_mult_2";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1794
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1795
 Goal "integ_of (v BIT b) zmod integ_of (w BIT False) = \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1796
 \         (if b then \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1797
 \               if #0 $<= integ_of w \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1798
 \               then #2 $* (integ_of v zmod integ_of w) $+ #1    \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1799
 \               else #2 $* ((integ_of v $+ #1) zmod integ_of w) - #1  \
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1800
 \          else #2 $* (integ_of v zmod integ_of w))";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1801
 by (simp_tac (simpset_of Int.thy addsimps [zadd_assoc, integ_of_BIT]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1802
 by (asm_simp_tac (simpset()
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1803
		   delsimps bin_arith_extra_simps@bin_rel_simps
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1804
		   addsimps [zmod_zmult_zmult1,
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1805
			     pos_zmod_mult_2, lemma, neg_zmod_mult_2]) 1);
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1806
 qed "zmod_integ_of_BIT";
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1807
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1808
 Addsimps [zmod_integ_of_BIT];
01cbbf33779b the rest of integer division
paulson
parents: 10635
diff changeset
  1809
*)