src/ZF/Integ/Bin.thy
author wenzelm
Sat, 01 May 2004 22:08:57 +0200
changeset 14695 9c78044b99c3
parent 14511 73493236e97f
child 16417 9bc16273c2d4
permissions -rw-r--r--
improved Term.invent_names;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     1
(*  Title:      ZF/ex/Bin.thy
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     2
    ID:         $Id$
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     5
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     6
   The sign Pls stands for an infinite string of leading 0's.
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     7
   The sign Min stands for an infinite string of leading 1's.
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     8
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     9
A number can have multiple representations, namely leading 0's with sign
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    10
Pls and leading 1's with sign Min.  See twos-compl.ML/int_of_binary for
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    11
the numerical interpretation.
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    12
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    13
The representation expects that (m mod 2) is 0 or 1, even if m is negative;
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    14
For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    15
*)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    16
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    17
header{*Arithmetic on Binary Integers*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    18
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    19
theory Bin = Int + Datatype:
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    20
6117
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 6046
diff changeset
    21
consts  bin :: i
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 6046
diff changeset
    22
datatype
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 6046
diff changeset
    23
  "bin" = Pls
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 6046
diff changeset
    24
        | Min
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
    25
        | Bit ("w: bin", "b: bool")	(infixl "BIT" 90)
6117
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 6046
diff changeset
    26
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 6046
diff changeset
    27
syntax
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    28
  "_Int"    :: "xnum => i"        ("_")
6117
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 6046
diff changeset
    29
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    30
consts
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    31
  integ_of  :: "i=>i"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    32
  NCons     :: "[i,i]=>i"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    33
  bin_succ  :: "i=>i"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    34
  bin_pred  :: "i=>i"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    35
  bin_minus :: "i=>i"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    36
  bin_adder :: "i=>i"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    37
  bin_mult  :: "[i,i]=>i"
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    38
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    39
primrec
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    40
  integ_of_Pls:  "integ_of (Pls)     = $# 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    41
  integ_of_Min:  "integ_of (Min)     = $-($#1)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    42
  integ_of_BIT:  "integ_of (w BIT b) = $#b $+ integ_of(w) $+ integ_of(w)"
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    43
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    44
    (** recall that cond(1,b,c)=b and cond(0,b,c)=0 **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    45
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    46
primrec (*NCons adds a bit, suppressing leading 0s and 1s*)
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    47
  NCons_Pls: "NCons (Pls,b)     = cond(b,Pls BIT b,Pls)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    48
  NCons_Min: "NCons (Min,b)     = cond(b,Min,Min BIT b)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    49
  NCons_BIT: "NCons (w BIT c,b) = w BIT c BIT b"
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    50
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
    51
primrec (*successor.  If a BIT, can change a 0 to a 1 without recursion.*)
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    52
  bin_succ_Pls:  "bin_succ (Pls)     = Pls BIT 1"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    53
  bin_succ_Min:  "bin_succ (Min)     = Pls"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    54
  bin_succ_BIT:  "bin_succ (w BIT b) = cond(b, bin_succ(w) BIT 0, NCons(w,1))"
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    55
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    56
primrec (*predecessor*)
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    57
  bin_pred_Pls:  "bin_pred (Pls)     = Min"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    58
  bin_pred_Min:  "bin_pred (Min)     = Min BIT 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    59
  bin_pred_BIT:  "bin_pred (w BIT b) = cond(b, NCons(w,0), bin_pred(w) BIT 1)"
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    60
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    61
primrec (*unary negation*)
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    62
  bin_minus_Pls:
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    63
    "bin_minus (Pls)       = Pls"
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    64
  bin_minus_Min:
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
    65
    "bin_minus (Min)       = Pls BIT 1"
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    66
  bin_minus_BIT:
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
    67
    "bin_minus (w BIT b) = cond(b, bin_pred(NCons(bin_minus(w),0)),
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
    68
				bin_minus(w) BIT 0)"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    69
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    70
primrec (*sum*)
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    71
  bin_adder_Pls:
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    72
    "bin_adder (Pls)     = (lam w:bin. w)"
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    73
  bin_adder_Min:
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    74
    "bin_adder (Min)     = (lam w:bin. bin_pred(w))"
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    75
  bin_adder_BIT:
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    76
    "bin_adder (v BIT x) = 
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    77
       (lam w:bin. 
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    78
         bin_case (v BIT x, bin_pred(v BIT x), 
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    79
                   %w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w),  
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    80
                               x xor y),
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    81
                   w))"
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    82
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    83
(*The bin_case above replaces the following mutually recursive function:
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    84
primrec
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
    85
  "adding (v,x,Pls)     = v BIT x"
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
    86
  "adding (v,x,Min)     = bin_pred(v BIT x)"
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    87
  "adding (v,x,w BIT y) = NCons(bin_adder (v, cond(x and y, bin_succ(w), w)), 
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
    88
				x xor y)"
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    89
*)
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    90
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    91
constdefs
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    92
  bin_add   :: "[i,i]=>i"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    93
    "bin_add(v,w) == bin_adder(v)`w"
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    94
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    95
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    96
primrec
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    97
  bin_mult_Pls:
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
    98
    "bin_mult (Pls,w)     = Pls"
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
    99
  bin_mult_Min:
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
   100
    "bin_mult (Min,w)     = bin_minus(w)"
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   101
  bin_mult_BIT:
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
   102
    "bin_mult (v BIT b,w) = cond(b, bin_add(NCons(bin_mult(v,w),0),w),
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6117
diff changeset
   103
				 NCons(bin_mult(v,w),0))"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   104
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   105
setup NumeralSyntax.setup
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   106
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   107
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   108
declare bin.intros [simp,TC]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   109
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   110
lemma NCons_Pls_0: "NCons(Pls,0) = Pls"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   111
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   112
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   113
lemma NCons_Pls_1: "NCons(Pls,1) = Pls BIT 1"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   114
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   115
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   116
lemma NCons_Min_0: "NCons(Min,0) = Min BIT 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   117
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   118
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   119
lemma NCons_Min_1: "NCons(Min,1) = Min"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   120
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   121
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   122
lemma NCons_BIT: "NCons(w BIT x,b) = w BIT x BIT b"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   123
by (simp add: bin.case_eqns)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   124
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   125
lemmas NCons_simps [simp] = 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   126
    NCons_Pls_0 NCons_Pls_1 NCons_Min_0 NCons_Min_1 NCons_BIT
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   127
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   128
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   129
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   130
(** Type checking **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   131
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   132
lemma integ_of_type [TC]: "w: bin ==> integ_of(w) : int"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   133
apply (induct_tac "w")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   134
apply (simp_all add: bool_into_nat)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   135
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   136
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   137
lemma NCons_type [TC]: "[| w: bin; b: bool |] ==> NCons(w,b) : bin"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   138
by (induct_tac "w", auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   139
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   140
lemma bin_succ_type [TC]: "w: bin ==> bin_succ(w) : bin"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   141
by (induct_tac "w", auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   142
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   143
lemma bin_pred_type [TC]: "w: bin ==> bin_pred(w) : bin"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   144
by (induct_tac "w", auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   145
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   146
lemma bin_minus_type [TC]: "w: bin ==> bin_minus(w) : bin"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   147
by (induct_tac "w", auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   148
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   149
(*This proof is complicated by the mutual recursion*)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   150
lemma bin_add_type [rule_format,TC]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   151
     "v: bin ==> ALL w: bin. bin_add(v,w) : bin"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   152
apply (unfold bin_add_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   153
apply (induct_tac "v")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   154
apply (rule_tac [3] ballI)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   155
apply (rename_tac [3] "w'")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   156
apply (induct_tac [3] "w'")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   157
apply (simp_all add: NCons_type)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   158
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   159
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   160
lemma bin_mult_type [TC]: "[| v: bin; w: bin |] ==> bin_mult(v,w) : bin"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   161
by (induct_tac "v", auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   162
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   163
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   164
subsubsection{*The Carry and Borrow Functions, 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   165
            @{term bin_succ} and @{term bin_pred}*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   166
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   167
(*NCons preserves the integer value of its argument*)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   168
lemma integ_of_NCons [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   169
     "[| w: bin; b: bool |] ==> integ_of(NCons(w,b)) = integ_of(w BIT b)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   170
apply (erule bin.cases)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   171
apply (auto elim!: boolE) 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   172
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   173
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   174
lemma integ_of_succ [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   175
     "w: bin ==> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   176
apply (erule bin.induct)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   177
apply (auto simp add: zadd_ac elim!: boolE) 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   178
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   179
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   180
lemma integ_of_pred [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   181
     "w: bin ==> integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   182
apply (erule bin.induct)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   183
apply (auto simp add: zadd_ac elim!: boolE) 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   184
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   185
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   186
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   187
subsubsection{*@{term bin_minus}: Unary Negation of Binary Integers*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   188
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   189
lemma integ_of_minus: "w: bin ==> integ_of(bin_minus(w)) = $- integ_of(w)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   190
apply (erule bin.induct)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   191
apply (auto simp add: zadd_ac zminus_zadd_distrib  elim!: boolE) 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   192
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   193
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   194
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   195
subsubsection{*@{term bin_add}: Binary Addition*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   196
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   197
lemma bin_add_Pls [simp]: "w: bin ==> bin_add(Pls,w) = w"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   198
by (unfold bin_add_def, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   199
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   200
lemma bin_add_Pls_right: "w: bin ==> bin_add(w,Pls) = w"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   201
apply (unfold bin_add_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   202
apply (erule bin.induct, auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   203
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   204
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   205
lemma bin_add_Min [simp]: "w: bin ==> bin_add(Min,w) = bin_pred(w)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   206
by (unfold bin_add_def, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   207
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   208
lemma bin_add_Min_right: "w: bin ==> bin_add(w,Min) = bin_pred(w)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   209
apply (unfold bin_add_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   210
apply (erule bin.induct, auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   211
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   212
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   213
lemma bin_add_BIT_Pls [simp]: "bin_add(v BIT x,Pls) = v BIT x"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   214
by (unfold bin_add_def, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   215
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   216
lemma bin_add_BIT_Min [simp]: "bin_add(v BIT x,Min) = bin_pred(v BIT x)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   217
by (unfold bin_add_def, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   218
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   219
lemma bin_add_BIT_BIT [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   220
     "[| w: bin;  y: bool |]               
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   221
      ==> bin_add(v BIT x, w BIT y) =  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   222
          NCons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   223
by (unfold bin_add_def, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   224
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   225
lemma integ_of_add [rule_format]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   226
     "v: bin ==>  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   227
          ALL w: bin. integ_of(bin_add(v,w)) = integ_of(v) $+ integ_of(w)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   228
apply (erule bin.induct, simp, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   229
apply (rule ballI)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   230
apply (induct_tac "wa")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   231
apply (auto simp add: zadd_ac elim!: boolE) 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   232
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   233
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   234
(*Subtraction*)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   235
lemma diff_integ_of_eq: 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   236
     "[| v: bin;  w: bin |]    
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   237
      ==> integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   238
apply (unfold zdiff_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   239
apply (simp add: integ_of_add integ_of_minus)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   240
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   241
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   242
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   243
subsubsection{*@{term bin_mult}: Binary Multiplication*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   244
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   245
lemma integ_of_mult:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   246
     "[| v: bin;  w: bin |]    
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   247
      ==> integ_of(bin_mult(v,w)) = integ_of(v) $* integ_of(w)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   248
apply (induct_tac "v", simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   249
apply (simp add: integ_of_minus)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   250
apply (auto simp add: zadd_ac integ_of_add zadd_zmult_distrib  elim!: boolE) 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   251
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   252
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   253
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   254
subsection{*Computations*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   255
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   256
(** extra rules for bin_succ, bin_pred **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   257
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   258
lemma bin_succ_1: "bin_succ(w BIT 1) = bin_succ(w) BIT 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   259
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   260
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   261
lemma bin_succ_0: "bin_succ(w BIT 0) = NCons(w,1)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   262
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   263
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   264
lemma bin_pred_1: "bin_pred(w BIT 1) = NCons(w,0)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   265
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   266
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   267
lemma bin_pred_0: "bin_pred(w BIT 0) = bin_pred(w) BIT 1"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   268
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   269
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   270
(** extra rules for bin_minus **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   271
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   272
lemma bin_minus_1: "bin_minus(w BIT 1) = bin_pred(NCons(bin_minus(w), 0))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   273
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   274
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   275
lemma bin_minus_0: "bin_minus(w BIT 0) = bin_minus(w) BIT 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   276
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   277
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   278
(** extra rules for bin_add **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   279
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   280
lemma bin_add_BIT_11: "w: bin ==> bin_add(v BIT 1, w BIT 1) =  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   281
                     NCons(bin_add(v, bin_succ(w)), 0)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   282
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   283
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   284
lemma bin_add_BIT_10: "w: bin ==> bin_add(v BIT 1, w BIT 0) =   
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   285
                     NCons(bin_add(v,w), 1)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   286
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   287
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   288
lemma bin_add_BIT_0: "[| w: bin;  y: bool |]  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   289
      ==> bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   290
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   291
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   292
(** extra rules for bin_mult **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   293
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   294
lemma bin_mult_1: "bin_mult(v BIT 1, w) = bin_add(NCons(bin_mult(v,w),0), w)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   295
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   296
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   297
lemma bin_mult_0: "bin_mult(v BIT 0, w) = NCons(bin_mult(v,w),0)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   298
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   299
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   300
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   301
(** Simplification rules with integer constants **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   302
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   303
lemma int_of_0: "$#0 = #0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   304
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   305
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   306
lemma int_of_succ: "$# succ(n) = #1 $+ $#n"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   307
by (simp add: int_of_add [symmetric] natify_succ)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   308
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   309
lemma zminus_0 [simp]: "$- #0 = #0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   310
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   311
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   312
lemma zadd_0_intify [simp]: "#0 $+ z = intify(z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   313
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   314
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   315
lemma zadd_0_right_intify [simp]: "z $+ #0 = intify(z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   316
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   317
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   318
lemma zmult_1_intify [simp]: "#1 $* z = intify(z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   319
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   320
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   321
lemma zmult_1_right_intify [simp]: "z $* #1 = intify(z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   322
by (subst zmult_commute, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   323
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   324
lemma zmult_0 [simp]: "#0 $* z = #0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   325
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   326
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   327
lemma zmult_0_right [simp]: "z $* #0 = #0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   328
by (subst zmult_commute, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   329
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   330
lemma zmult_minus1 [simp]: "#-1 $* z = $-z"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   331
by (simp add: zcompare_rls)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   332
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   333
lemma zmult_minus1_right [simp]: "z $* #-1 = $-z"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   334
apply (subst zmult_commute)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   335
apply (rule zmult_minus1)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   336
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   337
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   338
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   339
subsection{*Simplification Rules for Comparison of Binary Numbers*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   340
text{*Thanks to Norbert Voelker*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   341
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   342
(** Equals (=) **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   343
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   344
lemma eq_integ_of_eq: 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   345
     "[| v: bin;  w: bin |]    
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   346
      ==> ((integ_of(v)) = integ_of(w)) <->  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   347
          iszero (integ_of (bin_add (v, bin_minus(w))))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   348
apply (unfold iszero_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   349
apply (simp add: zcompare_rls integ_of_add integ_of_minus)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   350
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   351
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   352
lemma iszero_integ_of_Pls: "iszero (integ_of(Pls))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   353
by (unfold iszero_def, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   354
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   355
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   356
lemma nonzero_integ_of_Min: "~ iszero (integ_of(Min))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   357
apply (unfold iszero_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   358
apply (simp add: zminus_equation)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   359
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   360
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   361
lemma iszero_integ_of_BIT: 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   362
     "[| w: bin; x: bool |]  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   363
      ==> iszero (integ_of (w BIT x)) <-> (x=0 & iszero (integ_of(w)))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   364
apply (unfold iszero_def, simp)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   365
apply (subgoal_tac "integ_of (w) : int")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   366
apply typecheck
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   367
apply (drule int_cases)
13612
55d32e76ef4e Adapted to new simplifier.
berghofe
parents: 13560
diff changeset
   368
apply (safe elim!: boolE)
55d32e76ef4e Adapted to new simplifier.
berghofe
parents: 13560
diff changeset
   369
apply (simp_all (asm_lr) add: zcompare_rls zminus_zadd_distrib [symmetric]
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   370
                     int_of_add [symmetric])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   371
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   372
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   373
lemma iszero_integ_of_0:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   374
     "w: bin ==> iszero (integ_of (w BIT 0)) <-> iszero (integ_of(w))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   375
by (simp only: iszero_integ_of_BIT, blast) 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   376
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   377
lemma iszero_integ_of_1: "w: bin ==> ~ iszero (integ_of (w BIT 1))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   378
by (simp only: iszero_integ_of_BIT, blast)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   379
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   380
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   381
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   382
(** Less-than (<) **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   383
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   384
lemma less_integ_of_eq_neg: 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   385
     "[| v: bin;  w: bin |]    
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   386
      ==> integ_of(v) $< integ_of(w)  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   387
          <-> znegative (integ_of (bin_add (v, bin_minus(w))))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   388
apply (unfold zless_def zdiff_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   389
apply (simp add: integ_of_minus integ_of_add)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   390
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   391
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   392
lemma not_neg_integ_of_Pls: "~ znegative (integ_of(Pls))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   393
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   394
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   395
lemma neg_integ_of_Min: "znegative (integ_of(Min))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   396
by simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   397
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   398
lemma neg_integ_of_BIT:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   399
     "[| w: bin; x: bool |]  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   400
      ==> znegative (integ_of (w BIT x)) <-> znegative (integ_of(w))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   401
apply simp
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   402
apply (subgoal_tac "integ_of (w) : int")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   403
apply typecheck
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   404
apply (drule int_cases)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   405
apply (auto elim!: boolE simp add: int_of_add [symmetric]  zcompare_rls)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   406
apply (simp_all add: zminus_zadd_distrib [symmetric] zdiff_def 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   407
                     int_of_add [symmetric])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   408
apply (subgoal_tac "$#1 $- $# succ (succ (n #+ n)) = $- $# succ (n #+ n) ")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   409
 apply (simp add: zdiff_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   410
apply (simp add: equation_zminus int_of_diff [symmetric])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   411
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   412
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   413
(** Less-than-or-equals (<=) **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   414
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   415
lemma le_integ_of_eq_not_less:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   416
     "(integ_of(x) $<= (integ_of(w))) <-> ~ (integ_of(w) $< (integ_of(x)))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   417
by (simp add: not_zless_iff_zle [THEN iff_sym])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   418
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   419
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   420
(*Delete the original rewrites, with their clumsy conditional expressions*)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   421
declare bin_succ_BIT [simp del] 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   422
        bin_pred_BIT [simp del] 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   423
        bin_minus_BIT [simp del]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   424
        NCons_Pls [simp del]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   425
        NCons_Min [simp del]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   426
        bin_adder_BIT [simp del]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   427
        bin_mult_BIT [simp del]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   428
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   429
(*Hide the binary representation of integer constants*)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   430
declare integ_of_Pls [simp del] integ_of_Min [simp del] integ_of_BIT [simp del]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   431
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   432
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   433
lemmas bin_arith_extra_simps =
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   434
     integ_of_add [symmetric]   
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   435
     integ_of_minus [symmetric] 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   436
     integ_of_mult [symmetric]  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   437
     bin_succ_1 bin_succ_0 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   438
     bin_pred_1 bin_pred_0 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   439
     bin_minus_1 bin_minus_0  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   440
     bin_add_Pls_right bin_add_Min_right
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   441
     bin_add_BIT_0 bin_add_BIT_10 bin_add_BIT_11
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   442
     diff_integ_of_eq
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   443
     bin_mult_1 bin_mult_0 NCons_simps
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   444
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   445
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   446
(*For making a minimal simpset, one must include these default simprules
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   447
  of thy.  Also include simp_thms, or at least (~False)=True*)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   448
lemmas bin_arith_simps =
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   449
     bin_pred_Pls bin_pred_Min
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   450
     bin_succ_Pls bin_succ_Min
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   451
     bin_add_Pls bin_add_Min
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   452
     bin_minus_Pls bin_minus_Min
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   453
     bin_mult_Pls bin_mult_Min 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   454
     bin_arith_extra_simps
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   455
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   456
(*Simplification of relational operations*)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   457
lemmas bin_rel_simps =
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   458
     eq_integ_of_eq iszero_integ_of_Pls nonzero_integ_of_Min
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   459
     iszero_integ_of_0 iszero_integ_of_1
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   460
     less_integ_of_eq_neg
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   461
     not_neg_integ_of_Pls neg_integ_of_Min neg_integ_of_BIT
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   462
     le_integ_of_eq_not_less
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   463
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   464
declare bin_arith_simps [simp]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   465
declare bin_rel_simps [simp]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   466
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   467
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   468
(** Simplification of arithmetic when nested to the right **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   469
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   470
lemma add_integ_of_left [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   471
     "[| v: bin;  w: bin |]    
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   472
      ==> integ_of(v) $+ (integ_of(w) $+ z) = (integ_of(bin_add(v,w)) $+ z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   473
by (simp add: zadd_assoc [symmetric])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   474
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   475
lemma mult_integ_of_left [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   476
     "[| v: bin;  w: bin |]    
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   477
      ==> integ_of(v) $* (integ_of(w) $* z) = (integ_of(bin_mult(v,w)) $* z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   478
by (simp add: zmult_assoc [symmetric])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   479
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   480
lemma add_integ_of_diff1 [simp]: 
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   481
    "[| v: bin;  w: bin |]    
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   482
      ==> integ_of(v) $+ (integ_of(w) $- c) = integ_of(bin_add(v,w)) $- (c)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   483
apply (unfold zdiff_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   484
apply (rule add_integ_of_left, auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   485
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   486
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   487
lemma add_integ_of_diff2 [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   488
     "[| v: bin;  w: bin |]    
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   489
      ==> integ_of(v) $+ (c $- integ_of(w)) =  
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   490
          integ_of (bin_add (v, bin_minus(w))) $+ (c)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   491
apply (subst diff_integ_of_eq [symmetric])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   492
apply (simp_all add: zdiff_def zadd_ac)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   493
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   494
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   495
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   496
(** More for integer constants **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   497
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   498
declare int_of_0 [simp] int_of_succ [simp]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   499
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   500
lemma zdiff0 [simp]: "#0 $- x = $-x"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   501
by (simp add: zdiff_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   502
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   503
lemma zdiff0_right [simp]: "x $- #0 = intify(x)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   504
by (simp add: zdiff_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   505
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   506
lemma zdiff_self [simp]: "x $- x = #0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   507
by (simp add: zdiff_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   508
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   509
lemma znegative_iff_zless_0: "k: int ==> znegative(k) <-> k $< #0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   510
by (simp add: zless_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   511
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   512
lemma zero_zless_imp_znegative_zminus: "[|#0 $< k; k: int|] ==> znegative($-k)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   513
by (simp add: zless_def)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   514
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   515
lemma zero_zle_int_of [simp]: "#0 $<= $# n"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   516
by (simp add: not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   517
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   518
lemma nat_of_0 [simp]: "nat_of(#0) = 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   519
by (simp only: natify_0 int_of_0 [symmetric] nat_of_int_of)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   520
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   521
lemma nat_le_int0_lemma: "[| z $<= $#0; z: int |] ==> nat_of(z) = 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   522
by (auto simp add: znegative_iff_zless_0 [THEN iff_sym] zle_def zneg_nat_of)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   523
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   524
lemma nat_le_int0: "z $<= $#0 ==> nat_of(z) = 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   525
apply (subgoal_tac "nat_of (intify (z)) = 0")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   526
apply (rule_tac [2] nat_le_int0_lemma, auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   527
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   528
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   529
lemma int_of_eq_0_imp_natify_eq_0: "$# n = #0 ==> natify(n) = 0"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   530
by (rule not_znegative_imp_zero, auto)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   531
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   532
lemma nat_of_zminus_int_of: "nat_of($- $# n) = 0"
14511
73493236e97f updated treatment of znegative and nat_of
paulson
parents: 13612
diff changeset
   533
by (simp add: nat_of_def int_of_def raw_nat_of zminus image_intrel_int)
13560
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   534
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   535
lemma int_of_nat_of: "#0 $<= z ==> $# nat_of(z) = intify(z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   536
apply (rule not_zneg_nat_of_intify)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   537
apply (simp add: znegative_iff_zless_0 not_zless_iff_zle)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   538
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   539
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   540
declare int_of_nat_of [simp] nat_of_zminus_int_of [simp]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   541
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   542
lemma int_of_nat_of_if: "$# nat_of(z) = (if #0 $<= z then intify(z) else #0)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   543
by (simp add: int_of_nat_of znegative_iff_zless_0 not_zle_iff_zless)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   544
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   545
lemma zless_nat_iff_int_zless: "[| m: nat; z: int |] ==> (m < nat_of(z)) <-> ($#m $< z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   546
apply (case_tac "znegative (z) ")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   547
apply (erule_tac [2] not_zneg_nat_of [THEN subst])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   548
apply (auto dest: zless_trans dest!: zero_zle_int_of [THEN zle_zless_trans]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   549
            simp add: znegative_iff_zless_0)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   550
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   551
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   552
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   553
(** nat_of and zless **)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   554
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   555
(*An alternative condition is  $#0 <= w  *)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   556
lemma zless_nat_conj_lemma: "$#0 $< z ==> (nat_of(w) < nat_of(z)) <-> (w $< z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   557
apply (rule iff_trans)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   558
apply (rule zless_int_of [THEN iff_sym])
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   559
apply (auto simp add: int_of_nat_of_if simp del: zless_int_of)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   560
apply (auto elim: zless_asym simp add: not_zle_iff_zless)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   561
apply (blast intro: zless_zle_trans)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   562
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   563
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   564
lemma zless_nat_conj: "(nat_of(w) < nat_of(z)) <-> ($#0 $< z & w $< z)"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   565
apply (case_tac "$#0 $< z")
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   566
apply (auto simp add: zless_nat_conj_lemma nat_le_int0 not_zless_iff_zle)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   567
done
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   568
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   569
(*This simprule cannot be added unless we can find a way to make eq_integ_of_eq
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   570
  unconditional!
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   571
  [The condition "True" is a hack to prevent looping.
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   572
    Conditional rewrite rules are tried after unconditional ones, so a rule
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   573
    like eq_nat_number_of will be tried first to eliminate #mm=#nn.]
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   574
  lemma integ_of_reorient [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   575
       "True ==> (integ_of(w) = x) <-> (x = integ_of(w))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   576
  by auto
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   577
*)
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   578
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   579
lemma integ_of_minus_reorient [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   580
     "(integ_of(w) = $- x) <-> ($- x = integ_of(w))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   581
by auto
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   582
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   583
lemma integ_of_add_reorient [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   584
     "(integ_of(w) = x $+ y) <-> (x $+ y = integ_of(w))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   585
by auto
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   586
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   587
lemma integ_of_diff_reorient [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   588
     "(integ_of(w) = x $- y) <-> (x $- y = integ_of(w))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   589
by auto
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   590
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   591
lemma integ_of_mult_reorient [simp]:
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   592
     "(integ_of(w) = x $* y) <-> (x $* y = integ_of(w))"
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   593
by auto
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   594
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   595
ML
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   596
{*
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   597
val bin_pred_Pls = thm "bin_pred_Pls";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   598
val bin_pred_Min = thm "bin_pred_Min";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   599
val bin_minus_Pls = thm "bin_minus_Pls";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   600
val bin_minus_Min = thm "bin_minus_Min";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   601
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   602
val NCons_Pls_0 = thm "NCons_Pls_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   603
val NCons_Pls_1 = thm "NCons_Pls_1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   604
val NCons_Min_0 = thm "NCons_Min_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   605
val NCons_Min_1 = thm "NCons_Min_1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   606
val NCons_BIT = thm "NCons_BIT";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   607
val NCons_simps = thms "NCons_simps";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   608
val integ_of_type = thm "integ_of_type";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   609
val NCons_type = thm "NCons_type";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   610
val bin_succ_type = thm "bin_succ_type";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   611
val bin_pred_type = thm "bin_pred_type";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   612
val bin_minus_type = thm "bin_minus_type";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   613
val bin_add_type = thm "bin_add_type";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   614
val bin_mult_type = thm "bin_mult_type";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   615
val integ_of_NCons = thm "integ_of_NCons";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   616
val integ_of_succ = thm "integ_of_succ";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   617
val integ_of_pred = thm "integ_of_pred";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   618
val integ_of_minus = thm "integ_of_minus";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   619
val bin_add_Pls = thm "bin_add_Pls";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   620
val bin_add_Pls_right = thm "bin_add_Pls_right";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   621
val bin_add_Min = thm "bin_add_Min";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   622
val bin_add_Min_right = thm "bin_add_Min_right";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   623
val bin_add_BIT_Pls = thm "bin_add_BIT_Pls";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   624
val bin_add_BIT_Min = thm "bin_add_BIT_Min";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   625
val bin_add_BIT_BIT = thm "bin_add_BIT_BIT";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   626
val integ_of_add = thm "integ_of_add";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   627
val diff_integ_of_eq = thm "diff_integ_of_eq";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   628
val integ_of_mult = thm "integ_of_mult";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   629
val bin_succ_1 = thm "bin_succ_1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   630
val bin_succ_0 = thm "bin_succ_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   631
val bin_pred_1 = thm "bin_pred_1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   632
val bin_pred_0 = thm "bin_pred_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   633
val bin_minus_1 = thm "bin_minus_1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   634
val bin_minus_0 = thm "bin_minus_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   635
val bin_add_BIT_11 = thm "bin_add_BIT_11";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   636
val bin_add_BIT_10 = thm "bin_add_BIT_10";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   637
val bin_add_BIT_0 = thm "bin_add_BIT_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   638
val bin_mult_1 = thm "bin_mult_1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   639
val bin_mult_0 = thm "bin_mult_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   640
val int_of_0 = thm "int_of_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   641
val int_of_succ = thm "int_of_succ";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   642
val zminus_0 = thm "zminus_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   643
val zadd_0_intify = thm "zadd_0_intify";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   644
val zadd_0_right_intify = thm "zadd_0_right_intify";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   645
val zmult_1_intify = thm "zmult_1_intify";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   646
val zmult_1_right_intify = thm "zmult_1_right_intify";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   647
val zmult_0 = thm "zmult_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   648
val zmult_0_right = thm "zmult_0_right";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   649
val zmult_minus1 = thm "zmult_minus1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   650
val zmult_minus1_right = thm "zmult_minus1_right";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   651
val eq_integ_of_eq = thm "eq_integ_of_eq";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   652
val iszero_integ_of_Pls = thm "iszero_integ_of_Pls";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   653
val nonzero_integ_of_Min = thm "nonzero_integ_of_Min";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   654
val iszero_integ_of_BIT = thm "iszero_integ_of_BIT";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   655
val iszero_integ_of_0 = thm "iszero_integ_of_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   656
val iszero_integ_of_1 = thm "iszero_integ_of_1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   657
val less_integ_of_eq_neg = thm "less_integ_of_eq_neg";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   658
val not_neg_integ_of_Pls = thm "not_neg_integ_of_Pls";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   659
val neg_integ_of_Min = thm "neg_integ_of_Min";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   660
val neg_integ_of_BIT = thm "neg_integ_of_BIT";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   661
val le_integ_of_eq_not_less = thm "le_integ_of_eq_not_less";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   662
val bin_arith_extra_simps = thms "bin_arith_extra_simps";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   663
val bin_arith_simps = thms "bin_arith_simps";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   664
val bin_rel_simps = thms "bin_rel_simps";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   665
val add_integ_of_left = thm "add_integ_of_left";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   666
val mult_integ_of_left = thm "mult_integ_of_left";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   667
val add_integ_of_diff1 = thm "add_integ_of_diff1";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   668
val add_integ_of_diff2 = thm "add_integ_of_diff2";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   669
val zdiff0 = thm "zdiff0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   670
val zdiff0_right = thm "zdiff0_right";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   671
val zdiff_self = thm "zdiff_self";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   672
val znegative_iff_zless_0 = thm "znegative_iff_zless_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   673
val zero_zless_imp_znegative_zminus = thm "zero_zless_imp_znegative_zminus";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   674
val zero_zle_int_of = thm "zero_zle_int_of";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   675
val nat_of_0 = thm "nat_of_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   676
val nat_le_int0 = thm "nat_le_int0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   677
val int_of_eq_0_imp_natify_eq_0 = thm "int_of_eq_0_imp_natify_eq_0";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   678
val nat_of_zminus_int_of = thm "nat_of_zminus_int_of";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   679
val int_of_nat_of = thm "int_of_nat_of";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   680
val int_of_nat_of_if = thm "int_of_nat_of_if";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   681
val zless_nat_iff_int_zless = thm "zless_nat_iff_int_zless";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   682
val zless_nat_conj = thm "zless_nat_conj";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   683
val integ_of_minus_reorient = thm "integ_of_minus_reorient";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   684
val integ_of_add_reorient = thm "integ_of_add_reorient";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   685
val integ_of_diff_reorient = thm "integ_of_diff_reorient";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   686
val integ_of_mult_reorient = thm "integ_of_mult_reorient";
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   687
*}
d9651081578b conversion of ZF/Integ/{Int,Bin} to Isar scripts
paulson
parents: 12182
diff changeset
   688
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   689
end