src/ZF/Integ/Bin.ML
author wenzelm
Fri, 04 Jun 1999 19:57:31 +0200
changeset 6779 2912aff958bd
parent 6153 bff90585cce5
child 9207 0c294bd701ea
permissions -rw-r--r--
Calculation.thy: Setup transitivity rules for calculational proofs.
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
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    32
val NCons_simps = [NCons_Pls_0, NCons_Pls_1, 
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    33
		   NCons_Min_0, NCons_Min_1,
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    34
		   NCons_BIT];
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    35
Addsimps NCons_simps;
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    36
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    37
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    38
(** Type checking **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    39
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    40
Goal "w: bin ==> integ_of(w) : int";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    41
by (induct_tac "w" 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    42
by (ALLGOALS (asm_simp_tac (simpset() addsimps [bool_into_nat])));
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    43
qed "integ_of_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    44
AddTCs [integ_of_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    45
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    46
Goal "[| w: bin; b: bool |] ==> NCons(w,b) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    47
by (induct_tac "w" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    48
by Auto_tac;
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    49
qed "NCons_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    50
AddTCs [NCons_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    51
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    52
Goal "w: bin ==> bin_succ(w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    53
by (induct_tac "w" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    54
by Auto_tac;
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    55
qed "bin_succ_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    56
AddTCs [bin_succ_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    57
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    58
Goal "w: bin ==> bin_pred(w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    59
by (induct_tac "w" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    60
by Auto_tac;
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    61
qed "bin_pred_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    62
AddTCs [bin_pred_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    63
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    64
Goal "w: bin ==> bin_minus(w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    65
by (induct_tac "w" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    66
by Auto_tac;
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    67
qed "bin_minus_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    68
AddTCs [bin_minus_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    69
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    70
(*This proof is complicated by the mutual recursion*)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    71
Goal "v: bin ==> ALL w: bin. bin_add(v,w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    72
by (induct_tac "v" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    73
by (rtac ballI 3);
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    74
by (rename_tac "w'" 3);
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    75
by (induct_tac "w'" 3);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    76
by (ALLGOALS (asm_simp_tac (simpset() addsimps [NCons_type])));
6112
5e4871c5136b datatype package improvements
paulson
parents: 6065
diff changeset
    77
qed_spec_mp "bin_add_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    78
AddTCs [bin_add_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    79
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    80
Goal "[| v: bin; w: bin |] ==> bin_mult(v,w) : bin";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
    81
by (induct_tac "v" 1);
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    82
by Auto_tac;
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
    83
qed "bin_mult_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    84
AddTCs [bin_mult_type];
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    85
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    86
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    87
(**** The carry/borrow functions, bin_succ and bin_pred ****)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    88
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    89
(*NCons preserves the integer value of its argument*)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
    90
Goal "[| w: bin; b: bool |] ==>     \
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
    91
\         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
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   108
Goal "w: bin ==> integ_of(bin_pred(w)) = $~ ($#1) $+ integ_of(w)";
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
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   121
Goal "w: bin ==> integ_of(bin_minus(w)) = $~ integ_of(w)";
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
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   133
Goal "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
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   137
Goal "w: bin ==> bin_add(Min,w) = bin_pred(w)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   138
by (Asm_simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   139
qed "bin_add_Min";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   140
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   141
Goal "bin_add(v BIT x,Pls) = v BIT x";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   142
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   143
qed "bin_add_BIT_Pls";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   144
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   145
Goal "bin_add(v BIT x,Min) = bin_pred(v BIT x)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   146
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   147
qed "bin_add_BIT_Min";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   148
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   149
Goal "[| w: bin;  y: bool |]              \
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   150
\     ==> bin_add(v BIT x, w BIT y) = \
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   151
\         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
   152
by (Asm_simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   153
qed "bin_add_BIT_BIT";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   154
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   155
Addsimps [bin_add_Pls, bin_add_Min, bin_add_BIT_Pls,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   156
	  bin_add_BIT_Min, bin_add_BIT_BIT,
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   157
	  integ_of_succ, integ_of_pred];
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   158
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   159
Goal "v: bin ==> \
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   160
\         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
   161
by (etac bin.induct 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   162
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   163
by (Simp_tac 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   164
by (rtac ballI 1);
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
   165
by (induct_tac "wa" 1);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   166
by (ALLGOALS (asm_simp_tac (simpset() addsimps zadd_ac setloop (etac boolE))));
6112
5e4871c5136b datatype package improvements
paulson
parents: 6065
diff changeset
   167
qed_spec_mp "integ_of_add";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   168
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   169
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   170
(*** bin_add: binary multiplication ***)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   171
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   172
Goal "[| v: bin;  w: bin |]   \
2c8a8be36c94 converted to use new primrec section
paulson
parents: 5561
diff changeset
   173
\     ==> integ_of(bin_mult(v,w)) = integ_of(v) $* integ_of(w)";
6065
3b4a29166f26 induct_tac and exhaust_tac
paulson
parents: 6046
diff changeset
   174
by (induct_tac "v" 1);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   175
by (Asm_simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   176
by (asm_simp_tac (simpset() addsimps [integ_of_minus]) 1);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   177
by (etac boolE 1);
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   178
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib]) 2);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   179
by (asm_simp_tac (simpset() addsimps [integ_of_add,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   180
				      zadd_zmult_distrib] @ zadd_ac) 1);
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   181
qed "integ_of_mult";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   182
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   183
(**** Computations ****)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   184
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   185
(** extra rules for bin_succ, bin_pred **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   186
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   187
Goal "bin_succ(w BIT 1) = bin_succ(w) BIT 0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   188
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   189
qed "bin_succ_BIT1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   190
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   191
Goal "bin_succ(w BIT 0) = NCons(w,1)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   192
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   193
qed "bin_succ_BIT0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   194
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   195
Goal "bin_pred(w BIT 1) = NCons(w,0)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   196
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   197
qed "bin_pred_BIT1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   198
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   199
Goal "bin_pred(w BIT 0) = bin_pred(w) BIT 1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   200
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   201
qed "bin_pred_BIT0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   202
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   203
(** extra rules for bin_minus **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   204
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   205
Goal "bin_minus(w BIT 1) = bin_pred(NCons(bin_minus(w), 0))";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   206
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   207
qed "bin_minus_BIT1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   208
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   209
Goal "bin_minus(w BIT 0) = bin_minus(w) BIT 0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   210
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   211
qed "bin_minus_BIT0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   212
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   213
(** extra rules for bin_add **)
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   214
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   215
Goal "w: bin ==> bin_add(v BIT 1, w BIT 1) = \
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   216
\                    NCons(bin_add(v, bin_succ(w)), 0)";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   217
by (Asm_simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   218
qed "bin_add_BIT_BIT11";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   219
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   220
Goal "w: bin ==> bin_add(v BIT 1, w BIT 0) =  \
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   221
\                    NCons(bin_add(v,w), 1)";
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   222
by (Asm_simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   223
qed "bin_add_BIT_BIT10";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   224
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   225
Goal "[| w: bin;  y: bool |] ==> \
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   226
\           bin_add(v BIT 0, w BIT y) = NCons(bin_add(v,w), y)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   227
by (Asm_simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   228
qed "bin_add_BIT_BIT0";
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_mult **)
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 "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
   233
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   234
qed "bin_mult_BIT1";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   235
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   236
Goal "bin_mult(v BIT 0, w) = NCons(bin_mult(v,w),0)";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   237
by (Simp_tac 1);
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   238
qed "bin_mult_BIT0";
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   239
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   240
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   241
(*Delete the original rewrites, with their clumsy conditional expressions*)
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   242
Delsimps [bin_succ_BIT, bin_pred_BIT, bin_minus_BIT, 
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   243
          NCons_Pls, NCons_Min, bin_add_BIT, bin_mult_BIT];
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   244
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   245
(*Hide the binary representation of integer constants*)
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   246
Delsimps [integ_of_Pls, integ_of_Min, integ_of_BIT];
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   247
5528
4896b4e4077b new directory for Integers
paulson
parents:
diff changeset
   248
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   249
Addsimps [integ_of_add RS sym,   (*invoke bin_add*)
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   250
	  integ_of_minus RS sym, (*invoke bin_minus*)
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   251
	  integ_of_mult RS sym,  (*invoke bin_mult*)
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   252
	  bin_succ_Pls, bin_succ_Min,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   253
	  bin_succ_BIT1, bin_succ_BIT0,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   254
	  bin_pred_Pls, bin_pred_Min,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   255
	  bin_pred_BIT1, bin_pred_BIT0,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   256
	  bin_minus_Pls, bin_minus_Min,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   257
	  bin_minus_BIT1, bin_minus_BIT0,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   258
	  bin_add_Pls, bin_add_Min, bin_add_BIT_Pls, 
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   259
	  bin_add_BIT_Min, bin_add_BIT_BIT0, 
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   260
	  bin_add_BIT_BIT10, bin_add_BIT_BIT11,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   261
	  bin_mult_Pls, bin_mult_Min,
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6112
diff changeset
   262
	  bin_mult_BIT1, bin_mult_BIT0];