src/HOL/Integ/IntArith.thy
author paulson
Sun, 15 Feb 2004 10:46:37 +0100
changeset 14387 e96d5c42c4b0
parent 14378 69c4d5997669
child 14390 55fe71faadda
permissions -rw-r--r--
Polymorphic treatment of binary arithmetic using axclasses
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
     1
(*  Title:      HOL/Integ/IntArith.thy
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
     2
    ID:         $Id$
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
     3
    Authors:    Larry Paulson and Tobias Nipkow
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
     4
*)
12023
wenzelm
parents: 11868
diff changeset
     5
wenzelm
parents: 11868
diff changeset
     6
header {* Integer arithmetic *}
wenzelm
parents: 11868
diff changeset
     7
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 9214
diff changeset
     8
theory IntArith = Bin
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
     9
files ("int_arith1.ML"):
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 9214
diff changeset
    10
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    11
text{*Duplicate: can't understand why it's necessary*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    12
declare numeral_0_eq_0 [simp]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    13
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    14
subsection{*Instantiating Binary Arithmetic for the Integers*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    15
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    16
instance
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    17
  int :: number ..
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    18
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    19
primrec (*the type constraint is essential!*)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    20
  number_of_Pls: "number_of bin.Pls = 0"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    21
  number_of_Min: "number_of bin.Min = - (1::int)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    22
  number_of_BIT: "number_of(w BIT x) = (if x then 1 else 0) +
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    23
	                               (number_of w) + (number_of w)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    24
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    25
declare number_of_Pls [simp del]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    26
        number_of_Min [simp del]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    27
        number_of_BIT [simp del]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    28
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    29
instance int :: number_ring
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    30
proof
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    31
  show "Numeral0 = (0::int)" by (rule number_of_Pls)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    32
  show "-1 = - (1::int)" by (rule number_of_Min)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    33
  fix w :: bin and x :: bool
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    34
  show "(number_of (w BIT x) :: int) =
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    35
        (if x then 1 else 0) + number_of w + number_of w"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    36
    by (rule number_of_BIT)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    37
qed
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    38
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    39
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
    40
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
    41
subsection{*Inequality Reasoning for the Arithmetic Simproc*}
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
    42
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    43
lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    44
by (cut_tac w = 0 in zless_nat_conj, auto)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    45
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
    46
lemma zless_imp_add1_zle: "w<z ==> w + (1::int) \<le> z"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
    47
apply (rule eq_Abs_Integ [of z])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
    48
apply (rule eq_Abs_Integ [of w])
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
    49
apply (simp add: linorder_not_le [symmetric] zle int_def zadd One_int_def)
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
    50
done
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
    51
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
    52
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    53
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    54
lemma add_numeral_0: "Numeral0 + a = (a::'a::number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    55
by simp 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    56
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    57
lemma add_numeral_0_right: "a + Numeral0 = (a::'a::number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    58
by simp
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    59
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    60
lemma mult_numeral_1: "Numeral1 * a = (a::'a::number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    61
by simp 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    62
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    63
lemma mult_numeral_1_right: "a * Numeral1 = (a::'a::number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    64
by simp
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    65
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    66
text{*Theorem lists for the cancellation simprocs. The use of binary numerals
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    67
for 0 and 1 reduces the number of special cases.*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    68
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    69
lemmas add_0s = add_numeral_0 add_numeral_0_right
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    70
lemmas mult_1s = mult_numeral_1 mult_numeral_1_right 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    71
                 mult_minus1 mult_minus1_right
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    72
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    73
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    74
subsection{*Special Arithmetic Rules for Abstract 0 and 1*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    75
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    76
text{*Arithmetic computations are defined for binary literals, which leaves 0
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    77
and 1 as special cases. Addition already has rules for 0, but not 1.
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    78
Multiplication and unary minus already have rules for both 0 and 1.*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    79
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    80
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    81
lemma binop_eq: "[|f x y = g x y; x = x'; y = y'|] ==> f x' y' = g x' y'"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    82
by simp
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    83
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    84
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    85
lemmas add_number_of_eq = number_of_add [symmetric]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    86
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    87
text{*Allow 1 on either or both sides*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    88
lemma one_add_one_is_two: "1 + 1 = (2::'a::number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    89
by (simp del: numeral_1_eq_1 add: numeral_1_eq_1 [symmetric] add_number_of_eq)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    90
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    91
lemmas add_special =
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    92
    one_add_one_is_two
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    93
    binop_eq [of "op +", OF add_number_of_eq numeral_1_eq_1 refl, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    94
    binop_eq [of "op +", OF add_number_of_eq refl numeral_1_eq_1, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    95
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    96
text{*Allow 1 on either or both sides (1-1 already simplifies to 0)*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    97
lemmas diff_special =
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    98
    binop_eq [of "op -", OF diff_number_of_eq numeral_1_eq_1 refl, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
    99
    binop_eq [of "op -", OF diff_number_of_eq refl numeral_1_eq_1, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   100
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   101
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   102
lemmas eq_special =
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   103
    binop_eq [of "op =", OF eq_number_of_eq numeral_0_eq_0 refl, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   104
    binop_eq [of "op =", OF eq_number_of_eq numeral_1_eq_1 refl, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   105
    binop_eq [of "op =", OF eq_number_of_eq refl numeral_0_eq_0, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   106
    binop_eq [of "op =", OF eq_number_of_eq refl numeral_1_eq_1, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   107
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   108
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   109
lemmas less_special =
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   110
  binop_eq [of "op <", OF less_number_of_eq_neg numeral_0_eq_0 refl, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   111
  binop_eq [of "op <", OF less_number_of_eq_neg numeral_1_eq_1 refl, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   112
  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_0_eq_0, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   113
  binop_eq [of "op <", OF less_number_of_eq_neg refl numeral_1_eq_1, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   114
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   115
text{*Allow 0 or 1 on either side with a binary numeral on the other*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   116
lemmas le_special =
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   117
    binop_eq [of "op \<le>", OF le_number_of_eq numeral_0_eq_0 refl, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   118
    binop_eq [of "op \<le>", OF le_number_of_eq numeral_1_eq_1 refl, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   119
    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_0_eq_0, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   120
    binop_eq [of "op \<le>", OF le_number_of_eq refl numeral_1_eq_1, standard]
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   121
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   122
lemmas arith_special = 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   123
       add_special diff_special eq_special less_special le_special
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   124
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   125
12023
wenzelm
parents: 11868
diff changeset
   126
use "int_arith1.ML"
wenzelm
parents: 11868
diff changeset
   127
setup int_arith_setup
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   128
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   129
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   130
subsection{*Lemmas About Small Numerals*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   131
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   132
lemma of_int_m1 [simp]: "of_int -1 = (-1 :: 'a :: number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   133
proof -
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   134
  have "(of_int -1 :: 'a) = of_int (- 1)" by simp
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   135
  also have "... = - of_int 1" by (simp only: of_int_minus)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   136
  also have "... = -1" by simp
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   137
  finally show ?thesis .
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   138
qed
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   139
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   140
lemma abs_minus_one [simp]: "abs (-1) = (1::'a::{ordered_ring,number_ring})"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   141
by (simp add: abs_if)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   142
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   143
lemma of_int_number_of_eq:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   144
     "of_int (number_of v) = (number_of v :: 'a :: number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   145
apply (induct v)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   146
apply (simp_all only: number_of of_int_add, simp_all) 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   147
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   148
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   149
text{*Lemmas for specialist use, NOT as default simprules*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   150
lemma mult_2: "2 * z = (z+z::'a::number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   151
proof -
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   152
  have "2*z = (1 + 1)*z" by simp
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   153
  also have "... = z+z" by (simp add: left_distrib)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   154
  finally show ?thesis .
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   155
qed
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   156
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   157
lemma mult_2_right: "z * 2 = (z+z::'a::number_ring)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   158
by (subst mult_commute, rule mult_2)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   159
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   160
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   161
subsection{*More Inequality Reasoning*}
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   162
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   163
lemma zless_add1_eq: "(w < z + (1::int)) = (w<z | w=z)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   164
by arith
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   165
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   166
lemma add1_zle_eq: "(w + (1::int) \<le> z) = (w<z)"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   167
by arith
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   168
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   169
lemma zle_diff1_eq [simp]: "(w \<le> z - (1::int)) = (w<(z::int))"
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   170
by arith
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   171
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   172
lemma zle_add1_eq_le [simp]: "(w < z + 1) = (w\<le>(z::int))"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   173
by arith
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   174
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   175
lemma int_one_le_iff_zero_less: "((1::int) \<le> z) = (0 < z)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   176
by arith
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   177
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   178
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   179
subsection{*The Functions @{term nat} and @{term int}*}
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   180
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   181
lemma nonneg_eq_int: "[| 0 \<le> z;  !!m. z = int m ==> P |] ==> P"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   182
by (blast dest: nat_0_le sym)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   183
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   184
lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   185
by auto
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   186
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   187
lemma nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   188
by auto
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   189
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   190
lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   191
apply (rule iffI)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   192
apply (erule nat_0_le [THEN subst])
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   193
apply (simp_all del: zless_int add: zless_int [symmetric]) 
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   194
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   195
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   196
lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   197
by (auto simp add: nat_eq_iff2)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   198
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   199
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   200
text{*Simplify the terms @{term "int 0"}, @{term "int(Suc 0)"} and
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   201
  @{term "w + - z"}*}
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   202
declare Zero_int_def [symmetric, simp]
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   203
declare One_int_def [symmetric, simp]
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   204
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   205
text{*cooper.ML refers to this theorem*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   206
lemmas zdiff_def_symmetric = zdiff_def [symmetric, simp]
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   207
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   208
lemma nat_0: "nat 0 = 0"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   209
by (simp add: nat_eq_iff)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   210
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   211
lemma nat_1: "nat 1 = Suc 0"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   212
by (subst nat_eq_iff, simp)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   213
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   214
lemma nat_2: "nat 2 = Suc (Suc 0)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   215
by (subst nat_eq_iff, simp)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   216
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   217
lemma nat_less_eq_zless: "0 \<le> w ==> (nat w < nat z) = (w<z)"
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   218
apply (case_tac "z < 0")
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   219
apply (auto simp add: nat_less_iff)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   220
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   221
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   222
lemma nat_le_eq_zle: "0 < w | 0 \<le> z ==> (nat w \<le> nat z) = (w\<le>z)"
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   223
by (auto simp add: linorder_not_less [symmetric] zless_nat_conj)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   224
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   225
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   226
text{*This simplifies expressions of the form @{term "int n = z"} where
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   227
      z is an integer literal.*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   228
declare int_eq_iff [of _ "number_of v", standard, simp]
13837
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13685
diff changeset
   229
13849
2584233cf3ef new simprule for int (nat n)
paulson
parents: 13837
diff changeset
   230
lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)"
2584233cf3ef new simprule for int (nat n)
paulson
parents: 13837
diff changeset
   231
  by simp
2584233cf3ef new simprule for int (nat n)
paulson
parents: 13837
diff changeset
   232
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14272
diff changeset
   233
lemma split_nat [arith_split]:
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   234
  "P(nat(i::int)) = ((\<forall>n. i = int n \<longrightarrow> P n) & (i < 0 \<longrightarrow> P 0))"
13575
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   235
  (is "?P = (?L & ?R)")
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   236
proof (cases "i < 0")
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   237
  case True thus ?thesis by simp
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   238
next
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   239
  case False
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   240
  have "?P = ?L"
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   241
  proof
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   242
    assume ?P thus ?L using False by clarsimp
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   243
  next
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   244
    assume ?L thus ?P using False by simp
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   245
  qed
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   246
  with False show ?thesis by simp
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   247
qed
ecb6ecd9af13 added nat_split
nipkow
parents: 12023
diff changeset
   248
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   249
13685
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   250
subsubsection "Induction principles for int"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   251
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   252
                     (* `set:int': dummy construction *)
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   253
theorem int_ge_induct[case_names base step,induct set:int]:
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   254
  assumes ge: "k \<le> (i::int)" and
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   255
        base: "P(k)" and
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   256
        step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   257
  shows "P i"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   258
proof -
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   259
  { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k \<le> i \<Longrightarrow> P i"
13685
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   260
    proof (induct n)
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   261
      case 0
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   262
      hence "i = k" by arith
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   263
      thus "P i" using base by simp
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   264
    next
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   265
      case (Suc n)
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   266
      hence "n = nat((i - 1) - k)" by arith
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   267
      moreover
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   268
      have ki1: "k \<le> i - 1" using Suc.prems by arith
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   269
      ultimately
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   270
      have "P(i - 1)" by(rule Suc.hyps)
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   271
      from step[OF ki1 this] show ?case by simp
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   272
    qed
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   273
  }
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   274
  from this ge show ?thesis by fast
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   275
qed
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   276
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   277
                     (* `set:int': dummy construction *)
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   278
theorem int_gr_induct[case_names base step,induct set:int]:
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   279
  assumes gr: "k < (i::int)" and
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   280
        base: "P(k+1)" and
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   281
        step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   282
  shows "P i"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   283
apply(rule int_ge_induct[of "k + 1"])
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   284
  using gr apply arith
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   285
 apply(rule base)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   286
apply (rule step, simp+)
13685
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   287
done
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   288
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   289
theorem int_le_induct[consumes 1,case_names base step]:
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   290
  assumes le: "i \<le> (k::int)" and
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   291
        base: "P(k)" and
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   292
        step: "\<And>i. \<lbrakk>i \<le> k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   293
  shows "P i"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   294
proof -
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14271
diff changeset
   295
  { fix n have "\<And>i::int. n = nat(k-i) \<Longrightarrow> i \<le> k \<Longrightarrow> P i"
13685
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   296
    proof (induct n)
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   297
      case 0
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   298
      hence "i = k" by arith
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   299
      thus "P i" using base by simp
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   300
    next
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   301
      case (Suc n)
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   302
      hence "n = nat(k - (i+1))" by arith
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   303
      moreover
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   304
      have ki1: "i + 1 \<le> k" using Suc.prems by arith
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   305
      ultimately
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   306
      have "P(i+1)" by(rule Suc.hyps)
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   307
      from step[OF ki1 this] show ?case by simp
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   308
    qed
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   309
  }
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   310
  from this le show ?thesis by fast
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   311
qed
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   312
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   313
theorem int_less_induct [consumes 1,case_names base step]:
13685
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   314
  assumes less: "(i::int) < k" and
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   315
        base: "P(k - 1)" and
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   316
        step: "\<And>i. \<lbrakk>i < k; P i\<rbrakk> \<Longrightarrow> P(i - 1)"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   317
  shows "P i"
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   318
apply(rule int_le_induct[of _ "k - 1"])
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   319
  using less apply arith
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   320
 apply(rule base)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   321
apply (rule step, simp+)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   322
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   323
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   324
subsection{*Intermediate value theorems*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   325
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   326
lemma int_val_lemma:
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   327
     "(\<forall>i<n::nat. abs(f(i+1) - f i) \<le> 1) -->  
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   328
      f 0 \<le> k --> k \<le> f n --> (\<exists>i \<le> n. f i = (k::int))"
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   329
apply (induct_tac "n", simp)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   330
apply (intro strip)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   331
apply (erule impE, simp)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   332
apply (erule_tac x = n in allE, simp)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   333
apply (case_tac "k = f (n+1) ")
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   334
 apply force
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   335
apply (erule impE)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   336
 apply (simp add: zabs_def split add: split_if_asm)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   337
apply (blast intro: le_SucI)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   338
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   339
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   340
lemmas nat0_intermed_int_val = int_val_lemma [rule_format (no_asm)]
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   341
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   342
lemma nat_intermed_int_val:
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   343
     "[| \<forall>i. m \<le> i & i < n --> abs(f(i + 1::nat) - f i) \<le> 1; m < n;  
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   344
         f m \<le> k; k \<le> f n |] ==> ? i. m \<le> i & i \<le> n & f i = (k::int)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   345
apply (cut_tac n = "n-m" and f = "%i. f (i+m) " and k = k 
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   346
       in int_val_lemma)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   347
apply simp
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   348
apply (erule impE)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   349
 apply (intro strip)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   350
 apply (erule_tac x = "i+m" in allE, arith)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   351
apply (erule exE)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   352
apply (rule_tac x = "i+m" in exI, arith)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   353
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   354
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   355
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   356
subsection{*Products and 1, by T. M. Rasmussen*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   357
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   358
lemma zmult_eq_self_iff: "(m = m*(n::int)) = (n = 1 | m = 0)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   359
apply auto
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   360
apply (subgoal_tac "m*1 = m*n")
14378
69c4d5997669 generic of_nat and of_int functions, and generalization of iszero
paulson
parents: 14365
diff changeset
   361
apply (drule mult_cancel_left [THEN iffD1], auto)
13685
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   362
done
0b8fbcf65d73 added induction thms
nipkow
parents: 13575
diff changeset
   363
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   364
text{*FIXME: tidy*}
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   365
lemma pos_zmult_eq_1_iff: "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   366
apply auto
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   367
apply (case_tac "m=1")
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   368
apply (case_tac [2] "n=1")
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   369
apply (case_tac [4] "m=1")
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   370
apply (case_tac [5] "n=1", auto)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   371
apply (tactic"distinct_subgoals_tac")
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   372
apply (subgoal_tac "1<m*n", simp)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   373
apply (rule less_1_mult, arith)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   374
apply (subgoal_tac "0<n", arith)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   375
apply (subgoal_tac "0<m*n")
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   376
apply (drule zero_less_mult_iff [THEN iffD1], auto)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   377
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   378
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   379
lemma zmult_eq_1_iff: "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   380
apply (case_tac "0<m")
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   381
apply (simp add: pos_zmult_eq_1_iff)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   382
apply (case_tac "m=0")
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   383
apply (simp del: number_of_reorient)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   384
apply (subgoal_tac "0 < -m")
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   385
apply (drule_tac n = "-n" in pos_zmult_eq_1_iff, auto)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   386
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   387
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   388
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   389
subsection{*More about nat*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   390
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   391
(*Analogous to zadd_int*)
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   392
lemma zdiff_int: "n \<le> m ==> int m - int n = int (m-n)"
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   393
by (induct m n rule: diff_induct, simp_all)
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   394
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   395
lemma nat_add_distrib:
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   396
     "[| (0::int) \<le> z;  0 \<le> z' |] ==> nat (z+z') = nat z + nat z'"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   397
apply (rule inj_int [THEN injD])
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   398
apply (simp add: zadd_int [symmetric])
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   399
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   400
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   401
lemma nat_diff_distrib:
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   402
     "[| (0::int) \<le> z';  z' \<le> z |] ==> nat (z-z') = nat z - nat z'"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   403
apply (rule inj_int [THEN injD])
14271
8ed6989228bb Simplification of the development of Integers
paulson
parents: 14266
diff changeset
   404
apply (simp add: zdiff_int [symmetric] nat_le_eq_zle)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   405
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   406
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   407
lemma nat_mult_distrib: "(0::int) \<le> z ==> nat (z*z') = nat z * nat z'"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   408
apply (case_tac "0 \<le> z'")
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   409
apply (rule inj_int [THEN injD])
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   410
apply (simp add: zmult_int [symmetric] zero_le_mult_iff)
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   411
apply (simp add: mult_le_0_iff)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   412
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   413
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   414
lemma nat_mult_distrib_neg: "z \<le> (0::int) ==> nat(z*z') = nat(-z) * nat(-z')"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   415
apply (rule trans)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   416
apply (rule_tac [2] nat_mult_distrib, auto)
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   417
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   418
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   419
lemma nat_abs_mult_distrib: "nat (abs (w * z)) = nat (abs w) * nat (abs z)"
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   420
apply (case_tac "z=0 | w=0")
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   421
apply (auto simp add: zabs_def nat_mult_distrib [symmetric] 
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   422
                      nat_mult_distrib_neg [symmetric] mult_less_0_iff)
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   423
done
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   424
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14295
diff changeset
   425
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   426
ML
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   427
{*
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   428
val zle_diff1_eq = thm "zle_diff1_eq";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   429
val zle_add1_eq_le = thm "zle_add1_eq_le";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   430
val nonneg_eq_int = thm "nonneg_eq_int";
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14378
diff changeset
   431
val abs_minus_one = thm "abs_minus_one";
14259
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   432
val nat_eq_iff = thm "nat_eq_iff";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   433
val nat_eq_iff2 = thm "nat_eq_iff2";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   434
val nat_less_iff = thm "nat_less_iff";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   435
val int_eq_iff = thm "int_eq_iff";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   436
val nat_0 = thm "nat_0";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   437
val nat_1 = thm "nat_1";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   438
val nat_2 = thm "nat_2";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   439
val nat_less_eq_zless = thm "nat_less_eq_zless";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   440
val nat_le_eq_zle = thm "nat_le_eq_zle";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   441
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   442
val nat_intermed_int_val = thm "nat_intermed_int_val";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   443
val zmult_eq_self_iff = thm "zmult_eq_self_iff";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   444
val pos_zmult_eq_1_iff = thm "pos_zmult_eq_1_iff";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   445
val zmult_eq_1_iff = thm "zmult_eq_1_iff";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   446
val nat_add_distrib = thm "nat_add_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   447
val nat_diff_distrib = thm "nat_diff_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   448
val nat_mult_distrib = thm "nat_mult_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   449
val nat_mult_distrib_neg = thm "nat_mult_distrib_neg";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   450
val nat_abs_mult_distrib = thm "nat_abs_mult_distrib";
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   451
*}
79f7d3451b1e conversion of ML to Isar scripts
paulson
parents: 13849
diff changeset
   452
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
   453
end