src/ZF/ex/Bin.ML
author lcp
Fri, 12 Aug 1994 12:28:46 +0200
changeset 515 abcc438e7c27
parent 477 53fc8ad84b33
child 760 f0200e91b272
permissions -rw-r--r--
installation of new inductive/datatype sections
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
     1
(*  Title: 	ZF/ex/Bin.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
     4
    Copyright   1994  University of Cambridge
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
     6
For Bin.thy.  Arithmetic on binary integers.
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
     9
open Bin;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
(*Perform induction on l, then prove the major premise using prems. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
fun bin_ind_tac a prems i = 
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    13
    EVERY [res_inst_tac [("x",a)] bin.induct i,
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
	   rename_last_tac a ["1"] (i+3),
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
	   ares_tac prems i];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    17
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    18
(** bin_rec -- by Vset recursion **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    19
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    20
goal Bin.thy "bin_rec(Plus,a,b,h) = a";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    21
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    22
by (rewrite_goals_tac bin.con_defs);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    23
by (simp_tac rank_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    24
val bin_rec_Plus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    25
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    26
goal Bin.thy "bin_rec(Minus,a,b,h) = b";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    27
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    28
by (rewrite_goals_tac bin.con_defs);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    29
by (simp_tac rank_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    30
val bin_rec_Minus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    31
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    32
goal Bin.thy "bin_rec(w$$x,a,b,h) = h(w, x, bin_rec(w,a,b,h))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    33
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    34
by (rewrite_goals_tac bin.con_defs);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    35
by (simp_tac rank_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    36
val bin_rec_Bcons = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    37
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    38
(*Type checking*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    39
val prems = goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    40
    "[| w: bin;    \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    41
\       a: C(Plus);   b: C(Minus);       \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    42
\       !!w x r. [| w: bin;  x: bool;  r: C(w) |] ==> h(w,x,r): C(w$$x)  \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    43
\    |] ==> bin_rec(w,a,b,h) : C(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    44
by (bin_ind_tac "w" prems 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    45
by (ALLGOALS 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    46
    (asm_simp_tac (ZF_ss addsimps (prems@[bin_rec_Plus,bin_rec_Minus,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    47
					 bin_rec_Bcons]))));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    48
val bin_rec_type = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    49
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    50
(** Versions for use with definitions **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    51
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    52
val [rew] = goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    53
    "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Plus) = a";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    54
by (rewtac rew);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    55
by (rtac bin_rec_Plus 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    56
val def_bin_rec_Plus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    57
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    58
val [rew] = goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    59
    "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Minus) = b";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    60
by (rewtac rew);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    61
by (rtac bin_rec_Minus 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    62
val def_bin_rec_Minus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    63
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    64
val [rew] = goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    65
    "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(w$$x) = h(w,x,j(w))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    66
by (rewtac rew);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    67
by (rtac bin_rec_Bcons 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    68
val def_bin_rec_Bcons = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    69
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    70
fun bin_recs def = map standard
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    71
	([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    72
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    73
(** Type checking **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    74
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    75
val bin_typechecks0 = bin_rec_type :: bin.intrs;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    76
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    77
goalw Bin.thy [integ_of_bin_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    78
    "!!w. w: bin ==> integ_of_bin(w) : integ";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    79
by (typechk_tac (bin_typechecks0@integ_typechecks@
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    80
		 nat_typechecks@[bool_into_nat]));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    81
val integ_of_bin_type = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    82
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    83
goalw Bin.thy [bin_succ_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    84
    "!!w. w: bin ==> bin_succ(w) : bin";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    85
by (typechk_tac (bin_typechecks0@bool_typechecks));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    86
val bin_succ_type = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    87
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    88
goalw Bin.thy [bin_pred_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    89
    "!!w. w: bin ==> bin_pred(w) : bin";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    90
by (typechk_tac (bin_typechecks0@bool_typechecks));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    91
val bin_pred_type = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    92
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    93
goalw Bin.thy [bin_minus_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    94
    "!!w. w: bin ==> bin_minus(w) : bin";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    95
by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    96
val bin_minus_type = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    97
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    98
goalw Bin.thy [bin_add_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    99
    "!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   100
by (typechk_tac ([bin_succ_type,bin_pred_type]@bin_typechecks0@
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   101
		 bool_typechecks@ZF_typechecks));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   102
val bin_add_type = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   103
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   104
goalw Bin.thy [bin_mult_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   105
    "!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   106
by (typechk_tac ([bin_minus_type,bin_add_type]@bin_typechecks0@
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   107
		 bool_typechecks));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   108
val bin_mult_type = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   109
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   110
val bin_typechecks = bin_typechecks0 @
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   111
    [integ_of_bin_type, bin_succ_type, bin_pred_type, 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   112
     bin_minus_type, bin_add_type, bin_mult_type];
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   113
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   114
val bin_ss = integ_ss 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   115
    addsimps([bool_1I, bool_0I,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   116
	     bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @ 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   117
	     bin_recs integ_of_bin_def @ bool_simps @ bin_typechecks);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   118
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   119
val typechecks = bin_typechecks @ integ_typechecks @ nat_typechecks @
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   120
                 [bool_subset_nat RS subsetD];
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   121
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   122
(**** The carry/borrow functions, bin_succ and bin_pred ****)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   123
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   124
(** Lemmas **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   125
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   126
goal Integ.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   127
    "!!z v. [| z $+ v = z' $+ v';  \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   128
\       z: integ; z': integ;  v: integ; v': integ;  w: integ |]   \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   129
\    ==> z $+ (v $+ w) = z' $+ (v' $+ w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   130
by (asm_simp_tac (integ_ss addsimps ([zadd_assoc RS sym])) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   131
val zadd_assoc_cong = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   132
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   133
goal Integ.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   134
    "!!z v w. [| z: integ;  v: integ;  w: integ |]   \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   135
\    ==> z $+ (v $+ w) = v $+ (z $+ w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   136
by (REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   137
val zadd_assoc_swap = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   138
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   139
val zadd_cong = 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   140
    read_instantiate_sg (sign_of Integ.thy) [("t","op $+")] subst_context2;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   141
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   142
val zadd_kill = (refl RS zadd_cong);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   143
val zadd_assoc_swap_kill = zadd_kill RSN (4, zadd_assoc_swap RS trans);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   144
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   145
(*Pushes 'constants' of the form $#m to the right -- LOOPS if two!*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   146
val zadd_assoc_znat = standard (znat_type RS zadd_assoc_swap);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   147
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   148
goal Integ.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   149
    "!!z w. [| z: integ;  w: integ |]   \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   150
\    ==> w $+ (z $+ (w $+ z)) = w $+ (w $+ (z $+ z))";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   151
by (REPEAT (ares_tac [zadd_kill, zadd_assoc_swap] 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   152
val zadd_swap_pairs = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   153
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   154
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   155
val carry_ss = bin_ss addsimps 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   156
               (bin_recs bin_succ_def @ bin_recs bin_pred_def);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   157
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   158
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   159
    "!!w. w: bin ==> integ_of_bin(bin_succ(w)) = $#1 $+ integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   160
by (etac bin.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   161
by (simp_tac (carry_ss addsimps [zadd_0_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   162
by (simp_tac (carry_ss addsimps [zadd_zminus_inverse]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   163
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   164
by (ALLGOALS (asm_simp_tac (carry_ss addsimps [zadd_assoc])));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   165
by (REPEAT (ares_tac (zadd_swap_pairs::typechecks) 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   166
val integ_of_bin_succ = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   167
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   168
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   169
    "!!w. w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   170
by (etac bin.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   171
by (simp_tac (carry_ss addsimps [zadd_0_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   172
by (simp_tac (carry_ss addsimps [zadd_zminus_inverse]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   173
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   174
by (ALLGOALS 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   175
    (asm_simp_tac 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   176
     (carry_ss addsimps [zadd_assoc RS sym,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   177
			zadd_zminus_inverse, zadd_zminus_inverse2])));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   178
by (REPEAT (ares_tac ([zadd_commute, zadd_cong, refl]@typechecks) 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   179
val integ_of_bin_pred = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   180
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   181
(*These two results replace the definitions of bin_succ and bin_pred*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   182
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   183
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   184
(*** bin_minus: (unary!) negation of binary integers ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   185
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   186
val bin_minus_ss =
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   187
    bin_ss addsimps (bin_recs bin_minus_def @
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   188
		    [integ_of_bin_succ, integ_of_bin_pred]);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   189
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   190
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   191
    "!!w. w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   192
by (etac bin.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   193
by (simp_tac (bin_minus_ss addsimps [zminus_0]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   194
by (simp_tac (bin_minus_ss addsimps [zadd_0_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   195
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   196
by (ALLGOALS 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   197
    (asm_simp_tac (bin_minus_ss addsimps [zminus_zadd_distrib, zadd_assoc])));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   198
val integ_of_bin_minus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   199
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   200
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   201
(*** bin_add: binary addition ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   202
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   203
goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Plus,w) = w";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   204
by (asm_simp_tac bin_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   205
val bin_add_Plus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   206
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   207
goalw Bin.thy [bin_add_def] "!!w. w: bin ==> bin_add(Minus,w) = bin_pred(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   208
by (asm_simp_tac bin_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   209
val bin_add_Minus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   210
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   211
goalw Bin.thy [bin_add_def] "bin_add(v$$x,Plus) = v$$x";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   212
by (simp_tac bin_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   213
val bin_add_Bcons_Plus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   214
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   215
goalw Bin.thy [bin_add_def] "bin_add(v$$x,Minus) = bin_pred(v$$x)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   216
by (simp_tac bin_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   217
val bin_add_Bcons_Minus = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   218
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   219
goalw Bin.thy [bin_add_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   220
    "!!w y. [| w: bin;  y: bool |] ==> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   221
\           bin_add(v$$x, w$$y) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   222
\           bin_add(v, cond(x and y, bin_succ(w), w)) $$ (x xor y)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   223
by (asm_simp_tac bin_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   224
val bin_add_Bcons_Bcons = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   225
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   226
val bin_add_rews = [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   227
		    bin_add_Bcons_Minus, bin_add_Bcons_Bcons,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   228
		    integ_of_bin_succ, integ_of_bin_pred];
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   229
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   230
val bin_add_ss = bin_ss addsimps ([bool_subset_nat RS subsetD] @ bin_add_rews);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   231
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   232
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   233
    "!!v. v: bin ==> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   234
\         ALL w: bin. integ_of_bin(bin_add(v,w)) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   235
\                     integ_of_bin(v) $+ integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   236
by (etac bin.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   237
by (simp_tac bin_add_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   238
by (simp_tac bin_add_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   239
by (rtac ballI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   240
by (bin_ind_tac "wa" [] 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   241
by (asm_simp_tac (bin_add_ss addsimps [zadd_0_right]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   242
by (asm_simp_tac bin_add_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   243
by (REPEAT (ares_tac (zadd_commute::typechecks) 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   244
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   245
by (asm_simp_tac (bin_add_ss addsimps [zadd_assoc, zadd_swap_pairs]) 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   246
by (REPEAT (ares_tac ([refl, zadd_kill, zadd_assoc_swap_kill]@typechecks) 2));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   247
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   248
by (ALLGOALS (asm_simp_tac (bin_add_ss addsimps [zadd_assoc,zadd_swap_pairs])));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   249
by (REPEAT (ares_tac ([refl, zadd_kill, zadd_assoc_swap_kill RS sym]@
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   250
		      typechecks) 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   251
val integ_of_bin_add_lemma = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   252
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   253
val integ_of_bin_add = integ_of_bin_add_lemma RS bspec;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   254
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   255
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   256
(*** bin_add: binary multiplication ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   257
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   258
val bin_mult_ss =
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   259
    bin_ss addsimps (bin_recs bin_mult_def @ 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   260
		       [integ_of_bin_minus, integ_of_bin_add]);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   261
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   262
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   263
val major::prems = goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   264
    "[| v: bin; w: bin |] ==>	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   265
\    integ_of_bin(bin_mult(v,w)) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   266
\    integ_of_bin(v) $* integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   267
by (cut_facts_tac prems 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   268
by (bin_ind_tac "v" [major] 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   269
by (asm_simp_tac (bin_mult_ss addsimps [zmult_0]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   270
by (asm_simp_tac (bin_mult_ss addsimps [zmult_1,zmult_zminus]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   271
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   272
by (asm_simp_tac (bin_mult_ss addsimps [zadd_zmult_distrib]) 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   273
by (asm_simp_tac 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   274
    (bin_mult_ss addsimps [zadd_zmult_distrib, zmult_1, zadd_assoc]) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   275
by (REPEAT (ares_tac ([zadd_commute, zadd_assoc_swap_kill RS sym]@
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   276
		      typechecks) 1));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   277
val integ_of_bin_mult = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   278
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   279
(**** Computations ****)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   280
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   281
(** extra rules for bin_succ, bin_pred **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   282
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   283
val [bin_succ_Plus, bin_succ_Minus, _] = bin_recs bin_succ_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   284
val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   285
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   286
goal Bin.thy "bin_succ(w$$1) = bin_succ(w) $$ 0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   287
by (simp_tac carry_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   288
val bin_succ_Bcons1 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   289
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   290
goal Bin.thy "bin_succ(w$$0) = w$$1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   291
by (simp_tac carry_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   292
val bin_succ_Bcons0 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   293
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   294
goal Bin.thy "bin_pred(w$$1) = w$$0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   295
by (simp_tac carry_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   296
val bin_pred_Bcons1 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   297
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   298
goal Bin.thy "bin_pred(w$$0) = bin_pred(w) $$ 1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   299
by (simp_tac carry_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   300
val bin_pred_Bcons0 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   301
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   302
(** extra rules for bin_minus **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   303
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   304
val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   305
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   306
goal Bin.thy "bin_minus(w$$1) = bin_pred(bin_minus(w) $$ 0)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   307
by (simp_tac bin_minus_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   308
val bin_minus_Bcons1 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   309
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   310
goal Bin.thy "bin_minus(w$$0) = bin_minus(w) $$ 0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   311
by (simp_tac bin_minus_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   312
val bin_minus_Bcons0 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   313
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   314
(** extra rules for bin_add **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   315
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   316
goal Bin.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   317
    "!!w. w: bin ==> bin_add(v$$1, w$$1) = bin_add(v, bin_succ(w)) $$ 0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   318
by (asm_simp_tac bin_add_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   319
val bin_add_Bcons_Bcons11 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   320
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   321
goal Bin.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   322
    "!!w. w: bin ==> bin_add(v$$1, w$$0) = bin_add(v,w) $$ 1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   323
by (asm_simp_tac bin_add_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   324
val bin_add_Bcons_Bcons10 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   325
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   326
goal Bin.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   327
    "!!w y.[| w: bin;  y: bool |] ==> bin_add(v$$0, w$$y) = bin_add(v,w) $$ y";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   328
by (asm_simp_tac bin_add_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   329
val bin_add_Bcons_Bcons0 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   330
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   331
(** extra rules for bin_mult **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   332
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   333
val [bin_mult_Plus, bin_mult_Minus, _] = bin_recs bin_mult_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   334
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   335
goal Bin.thy "bin_mult(v$$1, w) = bin_add(bin_mult(v,w)$$0, w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   336
by (simp_tac bin_mult_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   337
val bin_mult_Bcons1 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   338
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   339
goal Bin.thy "bin_mult(v$$0, w) = bin_mult(v,w)$$0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   340
by (simp_tac bin_mult_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   341
val bin_mult_Bcons0 = result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   342
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   343
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   344
(*** The computation simpset ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   345
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   346
val bin_comp_ss = integ_ss 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   347
    addsimps [bin_succ_Plus, bin_succ_Minus,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   348
	     bin_succ_Bcons1, bin_succ_Bcons0,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   349
	     bin_pred_Plus, bin_pred_Minus,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   350
	     bin_pred_Bcons1, bin_pred_Bcons0,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   351
	     bin_minus_Plus, bin_minus_Minus,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   352
	     bin_minus_Bcons1, bin_minus_Bcons0,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   353
	     bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   354
	     bin_add_Bcons_Minus, bin_add_Bcons_Bcons0, 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   355
	     bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   356
	     bin_mult_Plus, bin_mult_Minus,
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   357
	     bin_mult_Bcons1, bin_mult_Bcons0]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   358
    setsolver (type_auto_tac ([bool_1I, bool_0I] @ bin_typechecks0));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   359
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   360
(*** Examples of performing binary arithmetic by simplification ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   361
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   362
proof_timing := true;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   363
(*All runtimes below are on a SPARCserver 10*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   364
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   365
(* 13+19 = 32 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   366
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   367
    "bin_add(Plus$$1$$1$$0$$1, Plus$$1$$0$$0$$1$$1) = Plus$$1$$0$$0$$0$$0$$0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   368
by (simp_tac bin_comp_ss 1);	(*0.6 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   369
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   370
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   371
bin_add(binary_of_int 13, binary_of_int 19);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   372
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   373
(* 1234+5678 = 6912 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   374
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   375
    "bin_add(Plus$$1$$0$$0$$1$$1$$0$$1$$0$$0$$1$$0, \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   376
\	     Plus$$1$$0$$1$$1$$0$$0$$0$$1$$0$$1$$1$$1$$0) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   377
\    Plus$$1$$1$$0$$1$$1$$0$$0$$0$$0$$0$$0$$0$$0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   378
by (simp_tac bin_comp_ss 1);	(*2.6 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   379
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   380
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   381
bin_add(binary_of_int 1234, binary_of_int 5678);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   382
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   383
(* 1359-2468 = ~1109 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   384
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   385
    "bin_add(Plus$$1$$0$$1$$0$$1$$0$$0$$1$$1$$1$$1,		\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   386
\	     Minus$$0$$1$$1$$0$$0$$1$$0$$1$$1$$1$$0$$0) = 	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   387
\    Minus$$1$$0$$1$$1$$1$$0$$1$$0$$1$$0$$1$$1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   388
by (simp_tac bin_comp_ss 1);	(*2.3 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   389
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   390
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   391
bin_add(binary_of_int 1359, binary_of_int ~2468);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   392
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   393
(* 93746-46375 = 47371 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   394
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   395
    "bin_add(Plus$$1$$0$$1$$1$$0$$1$$1$$1$$0$$0$$0$$1$$1$$0$$0$$1$$0, \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   396
\	     Minus$$0$$1$$0$$0$$1$$0$$1$$0$$1$$1$$0$$1$$1$$0$$0$$1) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   397
\    Plus$$0$$1$$0$$1$$1$$1$$0$$0$$1$$0$$0$$0$$0$$1$$0$$1$$1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   398
by (simp_tac bin_comp_ss 1);	(*3.9 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   399
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   400
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   401
bin_add(binary_of_int 93746, binary_of_int ~46375);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   402
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   403
(* negation of 65745 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   404
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   405
    "bin_minus(Plus$$1$$0$$0$$0$$0$$0$$0$$0$$0$$1$$1$$0$$1$$0$$0$$0$$1) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   406
\    Minus$$0$$1$$1$$1$$1$$1$$1$$1$$1$$0$$0$$1$$0$$1$$1$$1$$1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   407
by (simp_tac bin_comp_ss 1);	(*0.6 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   408
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   409
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   410
bin_minus(binary_of_int 65745);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   411
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   412
(* negation of ~54321 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   413
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   414
    "bin_minus(Minus$$0$$0$$1$$0$$1$$0$$1$$1$$1$$1$$0$$0$$1$$1$$1$$1) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   415
\    Plus$$0$$1$$1$$0$$1$$0$$1$$0$$0$$0$$0$$1$$1$$0$$0$$0$$1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   416
by (simp_tac bin_comp_ss 1);	(*0.7 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   417
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   418
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   419
bin_minus(binary_of_int ~54321);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   420
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   421
(* 13*19 = 247 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   422
goal Bin.thy "bin_mult(Plus$$1$$1$$0$$1, Plus$$1$$0$$0$$1$$1) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   423
\               Plus$$1$$1$$1$$1$$0$$1$$1$$1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   424
by (simp_tac bin_comp_ss 1);	(*1.5 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   425
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   426
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   427
bin_mult(binary_of_int 13, binary_of_int 19);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   428
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   429
(* ~84 * 51 = ~4284 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   430
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   431
    "bin_mult(Minus$$0$$1$$0$$1$$1$$0$$0, Plus$$1$$1$$0$$0$$1$$1) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   432
\    Minus$$0$$1$$1$$1$$1$$0$$1$$0$$0$$0$$1$$0$$0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   433
by (simp_tac bin_comp_ss 1);	(*2.6 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   434
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   435
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   436
bin_mult(binary_of_int ~84, binary_of_int 51);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   437
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   438
(* 255*255 = 65025;  the worst case for 8-bit operands *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   439
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   440
    "bin_mult(Plus$$1$$1$$1$$1$$1$$1$$1$$1, \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   441
\             Plus$$1$$1$$1$$1$$1$$1$$1$$1) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   442
\        Plus$$1$$1$$1$$1$$1$$1$$1$$0$$0$$0$$0$$0$$0$$0$$0$$1";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   443
by (simp_tac bin_comp_ss 1);	(*9.8 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   444
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   445
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   446
bin_mult(binary_of_int 255, binary_of_int 255);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   447
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   448
(* 1359 * ~2468 = ~3354012 *)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   449
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   450
    "bin_mult(Plus$$1$$0$$1$$0$$1$$0$$0$$1$$1$$1$$1, 		\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   451
\	      Minus$$0$$1$$1$$0$$0$$1$$0$$1$$1$$1$$0$$0) = 	\
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   452
\    Minus$$0$$0$$1$$1$$0$$0$$1$$1$$0$$1$$0$$0$$1$$0$$0$$1$$1$$0$$0$$1$$0$$0";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   453
by (simp_tac bin_comp_ss 1);	(*13.7 secs*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   454
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   455
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   456
bin_mult(binary_of_int 1359, binary_of_int ~2468);