src/ZF/Integ/Bin.ML
author wenzelm
Fri, 06 Oct 2000 17:35:58 +0200
changeset 10168 50be659d4222
parent 9909 111ccda5009b
child 10635 b115460e5c50
permissions -rw-r--r--
final tuning;
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.ML
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
Arithmetic on binary integers.
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     7
*)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
     8
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
     9
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    10
AddTCs bin.intrs;
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    11
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    12
Goal "NCons(Pls,0) = Pls";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    13
by (Asm_simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    14
qed "NCons_Pls_0";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    15
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    16
Goal "NCons(Pls,1) = Pls BIT 1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    17
by (Asm_simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    18
qed "NCons_Pls_1";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    19
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    20
Goal "NCons(Min,0) = Min BIT 0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    21
by (Asm_simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    22
qed "NCons_Min_0";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    23
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    24
Goal "NCons(Min,1) = Min";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    25
by (Asm_simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    26
qed "NCons_Min_1";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    27
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    28
Goal "NCons(w BIT x,b) = w BIT x BIT b";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    29
by (asm_simp_tac (simpset() addsimps bin.case_eqns) 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    30
qed "NCons_BIT";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    31
9909
111ccda5009b tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9883
diff changeset
    32
bind_thms ("NCons_simps",
111ccda5009b tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9883
diff changeset
    33
 [NCons_Pls_0, NCons_Pls_1,
111ccda5009b tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9883
diff changeset
    34
  NCons_Min_0, NCons_Min_1,
111ccda5009b tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9883
diff changeset
    35
  NCons_BIT]);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    36
Addsimps NCons_simps;
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    37
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    38
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    39
(** Type checking **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    40
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    41
Goal "w: bin ==> integ_of(w) : int";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    42
by (induct_tac "w" 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    43
by (ALLGOALS (asm_simp_tac (simpset() addsimps [bool_into_nat])));
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    44
qed "integ_of_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    45
AddTCs [integ_of_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    46
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    47
Goal "[| w: bin; b: bool |] ==> NCons(w,b) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    48
by (induct_tac "w" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    49
by Auto_tac;
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    50
qed "NCons_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    51
AddTCs [NCons_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    52
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    53
Goal "w: bin ==> bin_succ(w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    54
by (induct_tac "w" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    55
by Auto_tac;
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    56
qed "bin_succ_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    57
AddTCs [bin_succ_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    58
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    59
Goal "w: bin ==> bin_pred(w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    60
by (induct_tac "w" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    61
by Auto_tac;
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    62
qed "bin_pred_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    63
AddTCs [bin_pred_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    64
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    65
Goal "w: bin ==> bin_minus(w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    66
by (induct_tac "w" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    67
by Auto_tac;
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    68
qed "bin_minus_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    69
AddTCs [bin_minus_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    70
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    71
(*This proof is complicated by the mutual recursion*)
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
    72
Goalw [bin_add_def] "v: bin ==> ALL w: bin. bin_add(v,w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    73
by (induct_tac "v" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    74
by (rtac ballI 3);
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    75
by (rename_tac "w'" 3);
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    76
by (induct_tac "w'" 3);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    77
by (ALLGOALS (asm_simp_tac (simpset() addsimps [NCons_type])));
6112
5e4871c5136b datatype package improvements
paulson
parents: 6065
diff changeset
    78
qed_spec_mp "bin_add_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    79
AddTCs [bin_add_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    80
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    81
Goal "[| v: bin; w: bin |] ==> bin_mult(v,w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    82
by (induct_tac "v" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    83
by Auto_tac;
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    84
qed "bin_mult_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    85
AddTCs [bin_mult_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    86
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    87
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    88
(**** The carry/borrow functions, bin_succ and bin_pred ****)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    89
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    90
(*NCons preserves the integer value of its argument*)
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
    91
Goal "[| w: bin; b: bool |] ==> integ_of(NCons(w,b)) = integ_of(w BIT b)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    92
by (etac bin.elim 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    93
by (Asm_simp_tac 3);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    94
by (ALLGOALS (etac boolE));
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    95
by (ALLGOALS Asm_simp_tac);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    96
qed "integ_of_NCons";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    97
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    98
Addsimps [integ_of_NCons];
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    99
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   100
Goal "w: bin ==> integ_of(bin_succ(w)) = $#1 $+ integ_of(w)";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   101
by (etac bin.induct 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   102
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   103
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   104
by (etac boolE 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   105
by (ALLGOALS (asm_simp_tac (simpset() addsimps zadd_ac)));
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   106
qed "integ_of_succ";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   107
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   108
Goal "w: bin ==> integ_of(bin_pred(w)) = $- ($#1) $+ integ_of(w)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   109
by (etac bin.induct 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   110
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   111
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   112
by (etac boolE 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   113
by (ALLGOALS (asm_simp_tac (simpset() addsimps zadd_ac)));
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   114
qed "integ_of_pred";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   115
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   116
Addsimps [integ_of_succ, integ_of_pred];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   117
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   118
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   119
(*** bin_minus: (unary!) negation of binary integers ***)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   120
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   121
Goal "w: bin ==> integ_of(bin_minus(w)) = $- integ_of(w)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   122
by (etac bin.induct 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   123
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   124
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   125
by (etac boolE 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   126
by (ALLGOALS 
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   127
    (asm_simp_tac (simpset() addsimps zadd_ac@[zminus_zadd_distrib])));
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   128
qed "integ_of_minus";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   129
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   130
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   131
(*** bin_add: binary addition ***)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   132
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
   133
Goalw [bin_add_def] "w: bin ==> bin_add(Pls,w) = w";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   134
by (Asm_simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   135
qed "bin_add_Pls";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   136
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   137
Goalw [bin_add_def] "w: bin ==> bin_add(w,Pls) = w";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   138
by (etac bin.induct 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   139
by Auto_tac;
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   140
qed "bin_add_Pls_right";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   141
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
   142
Goalw [bin_add_def] "w: bin ==> bin_add(Min,w) = bin_pred(w)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   143
by (Asm_simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   144
qed "bin_add_Min";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   145
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   146
Goalw [bin_add_def] "w: bin ==> bin_add(w,Min) = bin_pred(w)";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   147
by (etac bin.induct 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   148
by Auto_tac;
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   149
qed "bin_add_Min_right";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   150
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
   151
Goalw [bin_add_def] "bin_add(v BIT x,Pls) = v BIT x";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   152
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   153
qed "bin_add_BIT_Pls";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   154
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
   155
Goalw [bin_add_def] "bin_add(v BIT x,Min) = bin_pred(v BIT x)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   156
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   157
qed "bin_add_BIT_Min";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   158
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
   159
Goalw [bin_add_def] "[| w: bin;  y: bool |]              \
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   160
\     ==> bin_add(v BIT x, w BIT y) = \
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   161
\         NCons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   162
by (Asm_simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   163
qed "bin_add_BIT_BIT";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   164
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   165
Addsimps [bin_add_Pls, bin_add_Min, bin_add_BIT_Pls,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   166
	  bin_add_BIT_Min, bin_add_BIT_BIT,
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   167
	  integ_of_succ, integ_of_pred];
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   168
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   169
Goal "v: bin ==> \
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   170
\         ALL w: bin. integ_of(bin_add(v,w)) = integ_of(v) $+ integ_of(w)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   171
by (etac bin.induct 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   172
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   173
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   174
by (rtac ballI 1);
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
   175
by (induct_tac "wa" 1);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   176
by (ALLGOALS (asm_simp_tac (simpset() addsimps zadd_ac setloop (etac boolE))));
6112
5e4871c5136b datatype package improvements
paulson
parents: 6065
diff changeset
   177
qed_spec_mp "integ_of_add";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   178
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   179
(*Subtraction*)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   180
Goalw [zdiff_def]
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   181
     "[| v: bin;  w: bin |]   \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   182
\     ==> integ_of(v) $- integ_of(w) = integ_of(bin_add (v, bin_minus(w)))";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   183
by (asm_simp_tac (simpset() addsimps [integ_of_add, integ_of_minus]) 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   184
qed "diff_integ_of_eq";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   185
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   186
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
   187
(*** bin_mult: binary multiplication ***)
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   188
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   189
Goal "[| v: bin;  w: bin |]   \
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   190
\     ==> integ_of(bin_mult(v,w)) = integ_of(v) $* integ_of(w)";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
   191
by (induct_tac "v" 1);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   192
by (Asm_simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   193
by (asm_simp_tac (simpset() addsimps [integ_of_minus]) 1);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   194
by (etac boolE 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   195
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib]) 2);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   196
by (asm_simp_tac (simpset() addsimps [integ_of_add,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   197
				      zadd_zmult_distrib] @ zadd_ac) 1);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   198
qed "integ_of_mult";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   199
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   200
(**** Computations ****)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   201
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   202
(** extra rules for bin_succ, bin_pred **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   203
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   204
Goal "bin_succ(w BIT 1) = bin_succ(w) BIT 0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   205
by (Simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   206
qed "bin_succ_1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   207
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   208
Goal "bin_succ(w BIT 0) = NCons(w,1)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   209
by (Simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   210
qed "bin_succ_0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   211
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   212
Goal "bin_pred(w BIT 1) = NCons(w,0)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   213
by (Simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   214
qed "bin_pred_1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   215
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   216
Goal "bin_pred(w BIT 0) = bin_pred(w) BIT 1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   217
by (Simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   218
qed "bin_pred_0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   219
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   220
(** extra rules for bin_minus **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   221
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   222
Goal "bin_minus(w BIT 1) = bin_pred(NCons(bin_minus(w), 0))";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   223
by (Simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   224
qed "bin_minus_1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   225
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   226
Goal "bin_minus(w BIT 0) = bin_minus(w) BIT 0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   227
by (Simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   228
qed "bin_minus_0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   229
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   230
(** extra rules for bin_add **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   231
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   232
Goal "w: bin ==> bin_add(v BIT 1, w BIT 1) = \
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   233
\                    NCons(bin_add(v, bin_succ(w)), 0)";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   234
by (Asm_simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   235
qed "bin_add_BIT_11";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   236
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   237
Goal "w: bin ==> bin_add(v BIT 1, w BIT 0) =  \
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   238
\                    NCons(bin_add(v,w), 1)";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   239
by (Asm_simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   240
qed "bin_add_BIT_10";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   241
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   242
Goal "[| w: bin;  y: bool |] \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   243
\     ==> bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   244
by (Asm_simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   245
qed "bin_add_BIT_0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   246
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   247
(** extra rules for bin_mult **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   248
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   249
Goal "bin_mult(v BIT 1, w) = bin_add(NCons(bin_mult(v,w),0), w)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   250
by (Simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   251
qed "bin_mult_1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   252
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   253
Goal "bin_mult(v BIT 0, w) = NCons(bin_mult(v,w),0)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   254
by (Simp_tac 1);
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   255
qed "bin_mult_0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   256
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   257
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   258
(** Simplification rules with integer constants **)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   259
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   260
Goal "$#0 = #0";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   261
by (Simp_tac 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   262
qed "int_of_0";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   263
9576
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   264
Goal "$# succ(n) = #1 $+ $#n";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   265
by (simp_tac (simpset() addsimps [int_of_add RS sym, natify_succ]) 1);
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   266
qed "int_of_succ";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   267
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   268
Goal "$- #0 = #0";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   269
by (Simp_tac 1);
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   270
qed "zminus_0";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   271
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   272
Addsimps [zminus_0];
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   273
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   274
Goal "#0 $+ z = intify(z)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   275
by (Simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   276
qed "zadd_0_intify";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   277
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   278
Goal "z $+ #0 = intify(z)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   279
by (Simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   280
qed "zadd_0_right_intify";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   281
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   282
Addsimps [zadd_0_intify, zadd_0_right_intify];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   283
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   284
Goal "#1 $* z = intify(z)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   285
by (Simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   286
qed "zmult_1_intify";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   287
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   288
Goal "z $* #1 = intify(z)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   289
by (stac zmult_commute 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   290
by (Simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   291
qed "zmult_1_right_intify";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   292
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   293
Addsimps [zmult_1_intify, zmult_1_right_intify];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   294
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   295
Goal "#0 $* z = #0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   296
by (Simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   297
qed "zmult_0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   298
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   299
Goal "z $* #0 = #0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   300
by (stac zmult_commute 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   301
by (Simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   302
qed "zmult_0_right";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   303
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   304
Addsimps [zmult_0, zmult_0_right];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   305
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   306
Goal "#-1 $* z = $-z";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   307
by (simp_tac (simpset() addsimps zcompare_rls) 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   308
qed "zmult_minus1";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   309
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   310
Goal "z $* #-1 = $-z";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   311
by (stac zmult_commute 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   312
by (rtac zmult_minus1 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   313
qed "zmult_minus1_right";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   314
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   315
Addsimps [zmult_minus1, zmult_minus1_right];
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   316
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   317
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   318
(*** Simplification rules for comparison of binary numbers (N Voelker) ***)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   319
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   320
(** Equals (=) **)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   321
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   322
Goalw [iszero_def]
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   323
     "[| v: bin;  w: bin |]   \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   324
\     ==> ((integ_of(v)) = integ_of(w)) <-> \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   325
\         iszero (integ_of (bin_add (v, bin_minus(w))))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   326
by (asm_simp_tac (simpset() addsimps
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   327
              (zcompare_rls @ [integ_of_add, integ_of_minus])) 1); 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   328
qed "eq_integ_of_eq"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   329
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   330
Goalw [iszero_def] "iszero (integ_of(Pls))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   331
by (Simp_tac 1); 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   332
qed "iszero_integ_of_Pls";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   333
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   334
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   335
Goalw [iszero_def] "~ iszero (integ_of(Min))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   336
by (simp_tac (simpset() addsimps [zminus_equation]) 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   337
qed "nonzero_integ_of_Min"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   338
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   339
Goalw [iszero_def]
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   340
     "[| w: bin; x: bool |] \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   341
\     ==> iszero (integ_of (w BIT x)) <-> (x=0 & iszero (integ_of(w)))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   342
by (Simp_tac 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   343
by (subgoal_tac "integ_of(w) : int" 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   344
by (Asm_simp_tac 2);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   345
by (dtac int_cases 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   346
by (auto_tac (claset() addSEs [boolE], 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   347
              simpset() addsimps [int_of_add RS sym]));  
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   348
by (ALLGOALS (asm_full_simp_tac 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   349
	      (simpset() addsimps zcompare_rls @ 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   350
			  [zminus_zadd_distrib RS sym, int_of_add RS sym])));
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   351
by (subgoal_tac "znegative ($- $# succ(x)) <-> znegative ($# succ(x))" 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   352
by (Asm_simp_tac 2);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   353
by (Full_simp_tac 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   354
qed "iszero_integ_of_BIT"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   355
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   356
Goal "w: bin ==> iszero (integ_of (w BIT 0)) <-> iszero (integ_of(w))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   357
by (asm_simp_tac (ZF_ss addsimps [iszero_integ_of_BIT]) 1); 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   358
qed "iszero_integ_of_0"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   359
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   360
Goal "w: bin ==> ~ iszero (integ_of (w BIT 1))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   361
by (asm_simp_tac (ZF_ss addsimps [iszero_integ_of_BIT]) 1); 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   362
qed "iszero_integ_of_1"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   363
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   364
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   365
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   366
(** Less-than (<) **)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   367
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   368
Goalw [zless_def,zdiff_def] 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   369
     "[| v: bin;  w: bin |]   \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   370
\     ==> integ_of(v) $< integ_of(w) \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   371
\         <-> znegative (integ_of (bin_add (v, bin_minus(w))))";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   372
by (asm_simp_tac (simpset() addsimps [integ_of_minus, integ_of_add]) 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   373
qed "less_integ_of_eq_neg"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   374
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   375
Goal "~ znegative (integ_of(Pls))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   376
by (Simp_tac 1); 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   377
qed "not_neg_integ_of_Pls"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   378
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   379
Goal "znegative (integ_of(Min))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   380
by (Simp_tac 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   381
qed "neg_integ_of_Min"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   382
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   383
Goal "[| w: bin; x: bool |] \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   384
\     ==> znegative (integ_of (w BIT x)) <-> znegative (integ_of(w))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   385
by (Asm_simp_tac 1); 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   386
by (subgoal_tac "integ_of(w) : int" 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   387
by (Asm_simp_tac 2);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   388
by (dtac int_cases 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   389
by (auto_tac (claset() addSEs [boolE], 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   390
              simpset() addsimps [int_of_add RS sym] @ zcompare_rls));  
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   391
by (ALLGOALS (asm_full_simp_tac 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   392
	      (simpset() addsimps [zminus_zadd_distrib RS sym, zdiff_def,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   393
				   int_of_add RS sym])));
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   394
by (subgoal_tac "$#1 $- $# succ(succ(x #+ x)) = $- $# succ(x #+ x)" 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   395
by (asm_full_simp_tac (simpset() addsimps [zdiff_def])1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   396
by (asm_simp_tac (simpset() addsimps [equation_zminus, int_of_diff RS sym])1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   397
qed "neg_integ_of_BIT"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   398
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   399
Goal "w: bin ==> iszero (integ_of (w BIT 0)) <-> iszero (integ_of(w))"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   400
by (asm_simp_tac (ZF_ss addsimps [iszero_integ_of_BIT]) 1); 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   401
qed "iszero_integ_of_0"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   402
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   403
(** Less-than-or-equals (<=) **)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   404
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   405
Goal "(integ_of(x) $<= (integ_of(w))) <-> ~ (integ_of(w) $< (integ_of(x)))";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   406
by (simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym]) 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   407
qed "le_integ_of_eq_not_less"; 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   408
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9207
diff changeset
   409
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   410
(*Delete the original rewrites, with their clumsy conditional expressions*)
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   411
Delsimps [bin_succ_BIT, bin_pred_BIT, bin_minus_BIT, 
9207
0c294bd701ea removed the mutual recursion from "bin_add"
paulson
parents: 6153
diff changeset
   412
          NCons_Pls, NCons_Min, bin_adder_BIT, bin_mult_BIT];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   413
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   414
(*Hide the binary representation of integer constants*)
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   415
Delsimps [integ_of_Pls, integ_of_Min, integ_of_BIT];
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   416
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   417
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   418
bind_thms ("NCons_simps", 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   419
	   [NCons_Pls_0, NCons_Pls_1, NCons_Min_0, NCons_Min_1, NCons_BIT]);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   420
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   421
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   422
bind_thms ("bin_arith_extra_simps",
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   423
    [integ_of_add RS sym,   (*invoke bin_add*)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   424
     integ_of_minus RS sym, (*invoke bin_minus*)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   425
     integ_of_mult RS sym,  (*invoke bin_mult*)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   426
     bin_succ_1, bin_succ_0, 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   427
     bin_pred_1, bin_pred_0, 
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   428
     bin_minus_1, bin_minus_0,  
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   429
     bin_add_Pls_right, bin_add_Min_right,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   430
     bin_add_BIT_0, bin_add_BIT_10, bin_add_BIT_11,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   431
     diff_integ_of_eq,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   432
     bin_mult_1, bin_mult_0] @ NCons_simps);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   433
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   434
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   435
(*For making a minimal simpset, one must include these default simprules
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   436
  of thy.  Also include simp_thms, or at least (~False)=True*)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   437
bind_thms ("bin_arith_simps",
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   438
    [bin_pred_Pls, bin_pred_Min,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   439
     bin_succ_Pls, bin_succ_Min,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   440
     bin_add_Pls, bin_add_Min,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   441
     bin_minus_Pls, bin_minus_Min,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   442
     bin_mult_Pls, bin_mult_Min] @ bin_arith_extra_simps);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   443
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   444
(*Simplification of relational operations*)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   445
bind_thms ("bin_rel_simps",
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   446
    [eq_integ_of_eq, iszero_integ_of_Pls, nonzero_integ_of_Min,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   447
     iszero_integ_of_0, iszero_integ_of_1,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   448
     less_integ_of_eq_neg,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   449
     not_neg_integ_of_Pls, neg_integ_of_Min, neg_integ_of_BIT,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   450
     le_integ_of_eq_not_less]);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   451
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   452
Addsimps bin_arith_simps;
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   453
Addsimps bin_rel_simps;
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   454
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   455
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   456
(** Simplification of arithmetic when nested to the right **)
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   457
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   458
Goal "[| v: bin;  w: bin |]   \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   459
\     ==> integ_of(v) $+ (integ_of(w) $+ z) = (integ_of(bin_add(v,w)) $+ z)";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   460
by (asm_simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   461
qed "add_integ_of_left";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   462
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   463
Goal "[| v: bin;  w: bin |]   \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   464
\     ==> integ_of(v) $* (integ_of(w) $* z) = (integ_of(bin_mult(v,w)) $* z)";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   465
by (asm_simp_tac (simpset() addsimps [zmult_assoc RS sym]) 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   466
qed "mult_integ_of_left";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   467
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   468
Goalw [zdiff_def]
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   469
    "[| v: bin;  w: bin |]   \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   470
\     ==> integ_of(v) $+ (integ_of(w) $- c) = integ_of(bin_add(v,w)) $- (c)";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   471
by (rtac add_integ_of_left 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   472
by Auto_tac;  
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   473
qed "add_integ_of_diff1";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   474
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   475
Goal "[| v: bin;  w: bin |]   \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   476
\     ==> integ_of(v) $+ (c $- integ_of(w)) = \
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   477
\         integ_of (bin_add (v, bin_minus(w))) $+ (c)";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   478
by (stac (diff_integ_of_eq RS sym) 1);
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   479
by (ALLGOALS (asm_simp_tac (simpset() addsimps zdiff_def::zadd_ac)));
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   480
qed "add_integ_of_diff2";
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   481
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   482
Addsimps [add_integ_of_left, mult_integ_of_left,
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 9548
diff changeset
   483
	  add_integ_of_diff1, add_integ_of_diff2]; 
9576
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   484
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   485
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   486
(** More for integer constants **)
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   487
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   488
Addsimps [int_of_0, int_of_succ];
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   489
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   490
Goal "#0 $- x = $-x";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   491
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   492
qed "zdiff0";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   493
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   494
Goal "x $- #0 = intify(x)";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   495
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   496
qed "zdiff0_right";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   497
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   498
Goal "x $- x = #0";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   499
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   500
qed "zdiff_self";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   501
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   502
Addsimps [zdiff0, zdiff0_right, zdiff_self];
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   503
9883
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   504
Goal "k: int ==> znegative(k) <-> k $< #0";
9576
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   505
by (asm_simp_tac (simpset() addsimps [zless_def]) 1);
9883
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   506
qed "znegative_iff_zless_0";
9576
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   507
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   508
Goal "[| #0 $< k; k: int |] ==> znegative($-k)";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   509
by (asm_full_simp_tac (simpset() addsimps [zless_def]) 1);
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   510
qed "zero_zless_imp_znegative_zminus";
3df14e0a3a51 interim working version: more improvements to the integers
paulson
parents: 9570
diff changeset
   511
9883
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   512
Goal "#0 $<= $# n";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   513
by (simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym, 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   514
                                  znegative_iff_zless_0 RS iff_sym]) 1); 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   515
qed "zero_zle_int_of";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   516
AddIffs [zero_zle_int_of];
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   517
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   518
Goal "nat_of(#0) = 0";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   519
by (simp_tac (ZF_ss addsimps [natify_0, int_of_0 RS sym, nat_of_int_of]) 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   520
qed "nat_of_0";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   521
Addsimps [nat_of_0];
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   522
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   523
Goal "[| z $<= $#0; z: int |] ==> nat_of(z) = 0"; 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   524
by (auto_tac (claset(), 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   525
	      simpset() addsimps [znegative_iff_zless_0 RS iff_sym, 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   526
				  zle_def, zneg_nat_of])); 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   527
val lemma = result();
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   528
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   529
Goal "z $<= $#0 ==> nat_of(z) = 0"; 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   530
by (subgoal_tac "nat_of(intify(z)) = 0" 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   531
by (rtac lemma 2);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   532
by Auto_tac;  
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   533
qed "nat_le_int0";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   534
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   535
Goal "$# n = #0 \\<Longrightarrow> natify(n) = 0";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   536
by (rtac not_znegative_imp_zero 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   537
by Auto_tac;  
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   538
qed "int_of_eq_0_imp_natify_eq_0";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   539
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   540
Goalw [nat_of_def, raw_nat_of_def] "nat_of($- $# n) = 0";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   541
by (auto_tac((claset() addSDs [not_znegative_imp_zero, natify_int_of_eq], 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   542
              simpset())     delIffs [int_of_eq]));
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   543
by (rtac the_equality 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   544
by Auto_tac;  
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   545
by (blast_tac (claset() addIs [int_of_eq_0_imp_natify_eq_0, sym]) 1); 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   546
qed "nat_of_zminus_int_of";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   547
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   548
Goal "#0 $<= z \\<Longrightarrow> $# nat_of(z) = intify(z)";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   549
by (rtac not_zneg_nat_of_intify 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   550
by (asm_simp_tac (simpset() addsimps [znegative_iff_zless_0, 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   551
                                      not_zless_iff_zle]) 1); 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   552
qed "int_of_nat_of";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   553
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   554
Addsimps [int_of_nat_of, nat_of_zminus_int_of];
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   555
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   556
Goal "$# nat_of(z) = (if #0 $<= z then intify(z) else #0)";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   557
by (simp_tac (simpset() addsimps [int_of_nat_of,
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   558
                              znegative_iff_zless_0, not_zle_iff_zless]) 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   559
qed "int_of_nat_of_if";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   560
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   561
Goal "[| m: nat; z: int |] ==> (m < nat_of(z)) <-> ($#m $< z)";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   562
by (case_tac "znegative(z)" 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   563
by (etac (not_zneg_nat_of RS subst) 2);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   564
by (auto_tac (claset() addDs [zless_trans]
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   565
		       addSDs [zero_zle_int_of RS zle_zless_trans], 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   566
	      simpset() addsimps [znegative_iff_zless_0])); 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   567
qed "zless_nat_iff_int_zless";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   568
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   569
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   570
(** nat_of and zless **)
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   571
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   572
(*An alternative condition is  $#0 <= w  *)
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   573
Goal "$#0 $< z ==> (nat_of(w) < nat_of(z)) <-> (w $< z)";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   574
by (rtac iff_trans 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   575
by (rtac (zless_int_of RS iff_sym) 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   576
by (auto_tac (claset(), 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   577
       simpset() addsimps [int_of_nat_of_if] delsimps [zless_int_of]));  
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   578
by (auto_tac (claset() addEs [zless_asym],
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   579
              simpset() addsimps [not_zle_iff_zless]));  
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   580
by (blast_tac (claset() addIs [zless_zle_trans]) 1); 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   581
val lemma = result();
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   582
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   583
Goal "(nat_of(w) < nat_of(z)) <-> ($#0 $< z & w $< z)";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   584
by (case_tac "$#0 $< z" 1);
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   585
by (auto_tac (claset(), 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   586
	      simpset() addsimps [lemma, nat_le_int0, not_zless_iff_zle])); 
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   587
qed "zless_nat_conj";
c1c8647af477 a number of new theorems
paulson
parents: 9578
diff changeset
   588