src/ZF/ex/Bin.ML
author clasohm
Tue, 30 Jan 1996 13:42:57 +0100
changeset 1461 6bcb44e4d6e5
parent 906 6cd9c397f36a
child 2469 b50b8c0eec01
permissions -rw-r--r--
expanded tabs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
     1
(*  Title:      ZF/ex/Bin.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
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,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    14
           rename_last_tac a ["1"] (i+3),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    15
           ares_tac prems i];
0
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);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    24
qed "bin_rec_Plus";
515
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);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    30
qed "bin_rec_Minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    31
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    32
goal Bin.thy "bin_rec(Bcons(w,x),a,b,h) = h(w, x, bin_rec(w,a,b,h))";
515
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);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    36
qed "bin_rec_Bcons";
515
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);       \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    42
\       !!w x r. [| w: bin; x: bool; r: C(w) |] ==> h(w,x,r): C(Bcons(w,x))  \
515
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 
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    46
    (asm_simp_tac (ZF_ss addsimps (prems@[bin_rec_Plus, bin_rec_Minus,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    47
                                          bin_rec_Bcons]))));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    48
qed "bin_rec_type";
515
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);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    56
qed "def_bin_rec_Plus";
515
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);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    62
qed "def_bin_rec_Minus";
515
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
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    65
    "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Bcons(w,x)) = h(w,x,j(w))";
515
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);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    68
qed "def_bin_rec_Bcons";
515
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
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    71
        ([def] RL [def_bin_rec_Plus, def_bin_rec_Minus, def_bin_rec_Bcons]);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    72
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    73
goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,0) = Plus";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    74
by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    75
qed "norm_Bcons_Plus_0";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    76
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    77
goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Plus,1) = Bcons(Plus,1)";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    78
by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    79
qed "norm_Bcons_Plus_1";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    80
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    81
goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Minus,0) = Bcons(Minus,0)";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    82
by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    83
qed "norm_Bcons_Minus_0";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    84
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    85
goalw Bin.thy [norm_Bcons_def] "norm_Bcons(Minus,1) = Minus";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    86
by (asm_simp_tac (ZF_ss addsimps (bin.case_eqns @ bool_simps)) 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    87
qed "norm_Bcons_Minus_1";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    88
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    89
goalw Bin.thy [norm_Bcons_def]
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    90
    "norm_Bcons(Bcons(w,x),b) = Bcons(Bcons(w,x),b)";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    91
by (asm_simp_tac (ZF_ss addsimps bin.case_eqns) 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    92
qed "norm_Bcons_Bcons";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    93
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    94
val norm_Bcons_simps = [norm_Bcons_Plus_0, norm_Bcons_Plus_1, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    95
                        norm_Bcons_Minus_0, norm_Bcons_Minus_1,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    96
                        norm_Bcons_Bcons];
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    97
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    98
(** Type checking **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    99
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   100
val bin_typechecks0 = bin_rec_type :: bin.intrs;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   101
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   102
goalw Bin.thy [integ_of_bin_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   103
    "!!w. w: bin ==> integ_of_bin(w) : integ";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   104
by (typechk_tac (bin_typechecks0@integ_typechecks@
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   105
                 nat_typechecks@[bool_into_nat]));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   106
qed "integ_of_bin_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   107
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   108
goalw Bin.thy [norm_Bcons_def]
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   109
    "!!w. [| w: bin; b: bool |] ==> norm_Bcons(w,b) : bin";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   110
by (etac bin.elim 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   111
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps bin.case_eqns)));
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   112
by (typechk_tac (bin_typechecks0@bool_typechecks));
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   113
qed "norm_Bcons_type";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   114
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   115
goalw Bin.thy [bin_succ_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   116
    "!!w. w: bin ==> bin_succ(w) : bin";
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   117
by (typechk_tac ([norm_Bcons_type]@bin_typechecks0@bool_typechecks));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   118
qed "bin_succ_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   119
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   120
goalw Bin.thy [bin_pred_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   121
    "!!w. w: bin ==> bin_pred(w) : bin";
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   122
by (typechk_tac ([norm_Bcons_type]@bin_typechecks0@bool_typechecks));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   123
qed "bin_pred_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   124
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   125
goalw Bin.thy [bin_minus_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   126
    "!!w. w: bin ==> bin_minus(w) : bin";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   127
by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   128
qed "bin_minus_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   129
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   130
goalw Bin.thy [bin_add_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   131
    "!!v w. [| v: bin; w: bin |] ==> bin_add(v,w) : bin";
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   132
by (typechk_tac ([norm_Bcons_type, bin_succ_type, bin_pred_type]@
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   133
                 bin_typechecks0@ bool_typechecks@ZF_typechecks));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   134
qed "bin_add_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   135
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   136
goalw Bin.thy [bin_mult_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   137
    "!!v w. [| v: bin; w: bin |] ==> bin_mult(v,w) : bin";
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   138
by (typechk_tac ([norm_Bcons_type, bin_minus_type, bin_add_type]@
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   139
                 bin_typechecks0@ bool_typechecks));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   140
qed "bin_mult_type";
515
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 bin_typechecks = bin_typechecks0 @
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   143
    [integ_of_bin_type, norm_Bcons_type, bin_succ_type, bin_pred_type, 
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   144
     bin_minus_type, bin_add_type, bin_mult_type];
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   145
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   146
val bin_ss = integ_ss 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   147
    addsimps([bool_1I, bool_0I,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   148
             bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @ 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   149
             bin_recs integ_of_bin_def @ bool_simps @ bin_typechecks);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   150
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   151
val typechecks = bin_typechecks @ integ_typechecks @ nat_typechecks @
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   152
                 [bool_subset_nat RS subsetD];
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
(**** The carry/borrow functions, bin_succ and bin_pred ****)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   155
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   156
(** Lemmas **)
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 Integ.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   159
    "!!z v. [| z $+ v = z' $+ v';  \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   160
\       z: integ; z': integ;  v: integ; v': integ;  w: integ |]   \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   161
\    ==> z $+ (v $+ w) = z' $+ (v' $+ w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   162
by (asm_simp_tac (integ_ss addsimps ([zadd_assoc RS sym])) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   163
qed "zadd_assoc_cong";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   164
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   165
goal Integ.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   166
    "!!z v w. [| z: integ;  v: integ;  w: integ |]   \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   167
\    ==> z $+ (v $+ w) = v $+ (z $+ w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   168
by (REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   169
qed "zadd_assoc_swap";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   170
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   171
(*Pushes 'constants' of the form $#m to the right -- LOOPS if two!*)
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   172
bind_thm ("zadd_assoc_znat", (znat_type RS zadd_assoc_swap));
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   173
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   174
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   175
val carry_ss = bin_ss addsimps 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   176
               (bin_recs bin_succ_def @ bin_recs bin_pred_def);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   177
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   178
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   179
(*norm_Bcons preserves the integer value of its argument*)
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   180
goal Bin.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   181
    "!!w. [| w: bin; b: bool |] ==>     \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   182
\         integ_of_bin(norm_Bcons(w,b)) = integ_of_bin(Bcons(w,b))";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   183
by (etac bin.elim 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   184
by (asm_simp_tac (ZF_ss addsimps norm_Bcons_simps) 3);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   185
by (ALLGOALS (etac boolE));
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   186
by (ALLGOALS (asm_simp_tac (bin_ss addsimps (norm_Bcons_simps))));
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   187
qed "integ_of_bin_norm_Bcons";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   188
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   189
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   190
    "!!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
   191
by (etac bin.induct 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   192
by (simp_tac carry_ss 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   193
by (simp_tac carry_ss 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   194
by (etac boolE 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   195
by (ALLGOALS
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   196
    (asm_simp_tac (carry_ss addsimps integ_of_bin_norm_Bcons::zadd_ac)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   197
qed "integ_of_bin_succ";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   198
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   199
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   200
    "!!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
   201
by (etac bin.induct 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   202
by (simp_tac carry_ss 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   203
by (simp_tac carry_ss 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   204
by (etac boolE 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   205
by (ALLGOALS
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   206
    (asm_simp_tac (carry_ss addsimps integ_of_bin_norm_Bcons::zadd_ac)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   207
qed "integ_of_bin_pred";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   208
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   209
(*These two results replace the definitions of bin_succ and bin_pred*)
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
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   212
(*** bin_minus: (unary!) negation of binary integers ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   213
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   214
val bin_minus_ss =
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   215
    bin_ss addsimps (bin_recs bin_minus_def @
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   216
                    [integ_of_bin_succ, integ_of_bin_pred]);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   217
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   218
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   219
    "!!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
   220
by (etac bin.induct 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   221
by (simp_tac bin_minus_ss 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   222
by (simp_tac bin_minus_ss 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   223
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   224
by (ALLGOALS 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   225
    (asm_simp_tac (bin_minus_ss addsimps [zminus_zadd_distrib, zadd_assoc])));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   226
qed "integ_of_bin_minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   227
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   228
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   229
(*** bin_add: binary addition ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   230
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   231
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
   232
by (asm_simp_tac bin_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   233
qed "bin_add_Plus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   234
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   235
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
   236
by (asm_simp_tac bin_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   237
qed "bin_add_Minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   238
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   239
goalw Bin.thy [bin_add_def] "bin_add(Bcons(v,x),Plus) = Bcons(v,x)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   240
by (simp_tac bin_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   241
qed "bin_add_Bcons_Plus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   242
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   243
goalw Bin.thy [bin_add_def] "bin_add(Bcons(v,x),Minus) = bin_pred(Bcons(v,x))";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   244
by (simp_tac bin_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   245
qed "bin_add_Bcons_Minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   246
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   247
goalw Bin.thy [bin_add_def]
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   248
    "!!w y. [| w: bin;  y: bool |] ==> \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   249
\           bin_add(Bcons(v,x), Bcons(w,y)) = \
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   250
\           norm_Bcons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   251
by (asm_simp_tac bin_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   252
qed "bin_add_Bcons_Bcons";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   253
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   254
val bin_add_simps = [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   255
                     bin_add_Bcons_Minus, bin_add_Bcons_Bcons,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   256
                     integ_of_bin_succ, integ_of_bin_pred,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   257
                     integ_of_bin_norm_Bcons];
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   258
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   259
val bin_add_ss = 
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   260
    bin_ss addsimps ([bool_subset_nat RS subsetD] @ bin_add_simps);
515
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
goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   263
    "!!v. v: bin ==> \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   264
\         ALL w: bin. integ_of_bin(bin_add(v,w)) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   265
\                     integ_of_bin(v) $+ integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   266
by (etac bin.induct 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   267
by (simp_tac bin_add_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   268
by (simp_tac bin_add_ss 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   269
by (rtac ballI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   270
by (bin_ind_tac "wa" [] 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   271
by (asm_simp_tac bin_add_ss 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   272
by (asm_simp_tac (bin_add_ss addsimps zadd_ac) 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   273
by (etac boolE 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   274
by (asm_simp_tac (bin_add_ss addsimps zadd_ac) 2);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   275
by (etac boolE 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   276
by (ALLGOALS (asm_simp_tac (bin_add_ss addsimps zadd_ac)));
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   277
val integ_of_bin_add_lemma = result();
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   278
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   279
bind_thm("integ_of_bin_add", integ_of_bin_add_lemma RS bspec);
515
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
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   282
(*** bin_add: binary multiplication ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   283
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   284
val bin_mult_ss =
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   285
    bin_ss addsimps (bin_recs bin_mult_def @ 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   286
                       [integ_of_bin_minus, integ_of_bin_add,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   287
                        integ_of_bin_norm_Bcons]);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   288
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   289
val major::prems = goal Bin.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   290
    "[| v: bin; w: bin |] ==>   \
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   291
\    integ_of_bin(bin_mult(v,w)) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   292
\    integ_of_bin(v) $* integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   293
by (cut_facts_tac prems 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   294
by (bin_ind_tac "v" [major] 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   295
by (asm_simp_tac bin_mult_ss 1);
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   296
by (asm_simp_tac bin_mult_ss 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   297
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   298
by (asm_simp_tac (bin_mult_ss addsimps [zadd_zmult_distrib]) 2);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   299
by (asm_simp_tac 
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   300
    (bin_mult_ss addsimps ([zadd_zmult_distrib, zmult_1] @ zadd_ac)) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   301
qed "integ_of_bin_mult";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   302
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   303
(**** Computations ****)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   304
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   305
(** extra rules for bin_succ, bin_pred **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   306
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   307
val [bin_succ_Plus, bin_succ_Minus, _] = bin_recs bin_succ_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   308
val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   309
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   310
goal Bin.thy "bin_succ(Bcons(w,1)) = Bcons(bin_succ(w), 0)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   311
by (simp_tac carry_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   312
qed "bin_succ_Bcons1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   313
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   314
goal Bin.thy "bin_succ(Bcons(w,0)) = norm_Bcons(w,1)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   315
by (simp_tac carry_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   316
qed "bin_succ_Bcons0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   317
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   318
goal Bin.thy "bin_pred(Bcons(w,1)) = norm_Bcons(w,0)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   319
by (simp_tac carry_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   320
qed "bin_pred_Bcons1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   321
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   322
goal Bin.thy "bin_pred(Bcons(w,0)) = Bcons(bin_pred(w), 1)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   323
by (simp_tac carry_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   324
qed "bin_pred_Bcons0";
515
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
(** extra rules for bin_minus **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   327
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   328
val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   329
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   330
goal Bin.thy "bin_minus(Bcons(w,1)) = bin_pred(Bcons(bin_minus(w), 0))";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   331
by (simp_tac bin_minus_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   332
qed "bin_minus_Bcons1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   333
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   334
goal Bin.thy "bin_minus(Bcons(w,0)) = Bcons(bin_minus(w), 0)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   335
by (simp_tac bin_minus_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   336
qed "bin_minus_Bcons0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   337
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   338
(** extra rules for bin_add **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   339
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   340
goal Bin.thy 
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   341
    "!!w. w: bin ==> bin_add(Bcons(v,1), Bcons(w,1)) = \
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   342
\                    norm_Bcons(bin_add(v, bin_succ(w)), 0)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   343
by (asm_simp_tac bin_add_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   344
qed "bin_add_Bcons_Bcons11";
515
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
goal Bin.thy 
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   347
    "!!w. w: bin ==> bin_add(Bcons(v,1), Bcons(w,0)) =  \
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   348
\                    norm_Bcons(bin_add(v,w), 1)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   349
by (asm_simp_tac bin_add_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   350
qed "bin_add_Bcons_Bcons10";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   351
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   352
goal Bin.thy 
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   353
    "!!w y. [| w: bin;  y: bool |] ==> \
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   354
\           bin_add(Bcons(v,0), Bcons(w,y)) = norm_Bcons(bin_add(v,w), y)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   355
by (asm_simp_tac bin_add_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   356
qed "bin_add_Bcons_Bcons0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   357
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   358
(** extra rules for bin_mult **)
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
val [bin_mult_Plus, bin_mult_Minus, _] = bin_recs bin_mult_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   361
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   362
goal Bin.thy
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   363
    "bin_mult(Bcons(v,1), w) = bin_add(norm_Bcons(bin_mult(v,w),0), w)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   364
by (simp_tac bin_mult_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   365
qed "bin_mult_Bcons1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   366
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   367
goal Bin.thy "bin_mult(Bcons(v,0), w) = norm_Bcons(bin_mult(v,w),0)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   368
by (simp_tac bin_mult_ss 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   369
qed "bin_mult_Bcons0";
515
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
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   372
(*** The computation simpset ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   373
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   374
val bin_comp_ss = integ_ss 
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   375
    addsimps [integ_of_bin_add RS sym,   (*invoke bin_add*)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   376
              integ_of_bin_minus RS sym, (*invoke bin_minus*)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   377
              integ_of_bin_mult RS sym,  (*invoke bin_mult*)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   378
              bin_succ_Plus, bin_succ_Minus,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   379
              bin_succ_Bcons1, bin_succ_Bcons0,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   380
              bin_pred_Plus, bin_pred_Minus,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   381
              bin_pred_Bcons1, bin_pred_Bcons0,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   382
              bin_minus_Plus, bin_minus_Minus,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   383
              bin_minus_Bcons1, bin_minus_Bcons0,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   384
              bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   385
              bin_add_Bcons_Minus, bin_add_Bcons_Bcons0, 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   386
              bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   387
              bin_mult_Plus, bin_mult_Minus,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   388
              bin_mult_Bcons1, bin_mult_Bcons0] @
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   389
             norm_Bcons_simps
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   390
    setsolver (type_auto_tac ([bool_1I, bool_0I] @ bin_typechecks0));
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   391
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   392
(*** Examples of performing binary arithmetic by simplification ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   393
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   394
proof_timing := true;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   395
(*All runtimes below are on a SPARCserver 10*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   396
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   397
goal Bin.thy "#13  $+  #19 = #32";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   398
by (simp_tac bin_comp_ss 1);    (*0.4 secs*)
515
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 13, binary_of_int 19);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   402
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   403
goal Bin.thy "#1234  $+  #5678 = #6912";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   404
by (simp_tac bin_comp_ss 1);    (*1.3 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   405
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   406
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   407
bin_add(binary_of_int 1234, binary_of_int 5678);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   408
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   409
goal Bin.thy "#1359  $+  #~2468 = #~1109";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   410
by (simp_tac bin_comp_ss 1);    (*1.2 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   411
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   412
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   413
bin_add(binary_of_int 1359, binary_of_int ~2468);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   414
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   415
goal Bin.thy "#93746  $+  #~46375 = #47371";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   416
by (simp_tac bin_comp_ss 1);    (*1.9 secs*)
515
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_add(binary_of_int 93746, binary_of_int ~46375);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   420
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   421
goal Bin.thy "$~ #65745 = #~65745";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   422
by (simp_tac bin_comp_ss 1);    (*0.4 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   423
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   424
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   425
bin_minus(binary_of_int 65745);
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
(* negation of ~54321 *)
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   428
goal Bin.thy "$~ #~54321 = #54321";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   429
by (simp_tac bin_comp_ss 1);    (*0.5 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   430
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   431
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   432
bin_minus(binary_of_int ~54321);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   433
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   434
goal Bin.thy "#13  $*  #19 = #247";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   435
by (simp_tac bin_comp_ss 1);    (*0.7 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   436
result();
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
bin_mult(binary_of_int 13, binary_of_int 19);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   439
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   440
goal Bin.thy "#~84  $*  #51 = #~4284";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   441
by (simp_tac bin_comp_ss 1);    (*1.3 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   442
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   443
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   444
bin_mult(binary_of_int ~84, binary_of_int 51);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   445
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   446
(*The worst case for 8-bit operands *)
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   447
goal Bin.thy "#255  $*  #255 = #65025";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   448
by (simp_tac bin_comp_ss 1);    (*4.3 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   449
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   450
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   451
bin_mult(binary_of_int 255, binary_of_int 255);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   452
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   453
goal Bin.thy "#1359  $*  #~2468 = #~3354012";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   454
by (simp_tac bin_comp_ss 1);    (*6.1 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   455
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   456
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   457
bin_mult(binary_of_int 1359, binary_of_int ~2468);