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