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