src/ZF/ex/Bin.ML
author paulson
Wed, 15 Jul 1998 14:13:18 +0200
changeset 5147 825877190618
parent 5137 60205b0de9b9
child 5512 4327eec06849
permissions -rw-r--r--
More tidying and removal of "\!\!... from Goal commands
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
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    11
Addsimps bin.case_eqns;
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    12
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
(*Perform induction on l, then prove the major premise using prems. *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
fun bin_ind_tac a prems i = 
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    15
    EVERY [res_inst_tac [("x",a)] bin.induct i,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    16
           rename_last_tac a ["1"] (i+3),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    17
           ares_tac prems i];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
515
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
(** bin_rec -- by Vset recursion **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    21
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
    22
Goal "bin_rec(Plus,a,b,h) = a";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    23
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    24
by (rewrite_goals_tac bin.con_defs);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    25
by (simp_tac rank_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    26
qed "bin_rec_Plus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    27
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
    28
Goal "bin_rec(Minus,a,b,h) = b";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    29
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    30
by (rewrite_goals_tac bin.con_defs);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    31
by (simp_tac rank_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    32
qed "bin_rec_Minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    33
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
    34
Goal "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
    35
by (rtac (bin_rec_def RS def_Vrec RS trans) 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    36
by (rewrite_goals_tac bin.con_defs);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    37
by (simp_tac rank_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
    38
qed "bin_rec_Bcons";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    39
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    40
(*Type checking*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    41
val prems = goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    42
    "[| w: bin;    \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    43
\       a: C(Plus);   b: C(Minus);       \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    44
\       !!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
    45
\    |] ==> bin_rec(w,a,b,h) : C(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    46
by (bin_ind_tac "w" prems 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    47
by (ALLGOALS 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
    48
    (asm_simp_tac (simpset() addsimps (prems@[bin_rec_Plus, bin_rec_Minus,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    49
                                          bin_rec_Bcons]))));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    50
qed "bin_rec_type";
515
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
(** Versions for use with definitions **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    53
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    54
val [rew] = goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    55
    "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Plus) = a";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    56
by (rewtac rew);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    57
by (rtac bin_rec_Plus 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    58
qed "def_bin_rec_Plus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    59
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    60
val [rew] = goal Bin.thy
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    61
    "[| !!w. j(w)==bin_rec(w,a,b,h) |] ==> j(Minus) = b";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    62
by (rewtac rew);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    63
by (rtac bin_rec_Minus 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    64
qed "def_bin_rec_Minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    65
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    66
val [rew] = goal Bin.thy
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    67
    "[| !!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
    68
by (rewtac rew);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    69
by (rtac bin_rec_Bcons 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
    70
qed "def_bin_rec_Bcons";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    71
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
    72
fun bin_recs def = map standard
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    73
        ([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
    74
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
    75
Goalw [norm_Bcons_def] "norm_Bcons(Plus,0) = Plus";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    76
by (Asm_simp_tac 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    77
qed "norm_Bcons_Plus_0";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    78
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
    79
Goalw [norm_Bcons_def] "norm_Bcons(Plus,1) = Bcons(Plus,1)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    80
by (Asm_simp_tac 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    81
qed "norm_Bcons_Plus_1";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    82
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
    83
Goalw [norm_Bcons_def] "norm_Bcons(Minus,0) = Bcons(Minus,0)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    84
by (Asm_simp_tac 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    85
qed "norm_Bcons_Minus_0";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    86
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
    87
Goalw [norm_Bcons_def] "norm_Bcons(Minus,1) = Minus";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    88
by (Asm_simp_tac 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    89
qed "norm_Bcons_Minus_1";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    90
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
    91
Goalw [norm_Bcons_def]
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    92
    "norm_Bcons(Bcons(w,x),b) = Bcons(Bcons(w,x),b)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
    93
by (asm_simp_tac (simpset() addsimps bin.case_eqns) 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    94
qed "norm_Bcons_Bcons";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    95
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    96
val norm_Bcons_simps = [norm_Bcons_Plus_0, norm_Bcons_Plus_1, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    97
                        norm_Bcons_Minus_0, norm_Bcons_Minus_1,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
    98
                        norm_Bcons_Bcons];
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
    99
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   100
(** Type checking **)
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
val bin_typechecks0 = bin_rec_type :: bin.intrs;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   103
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   104
Goalw [integ_of_bin_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   105
    "w: bin ==> integ_of_bin(w) : integ";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   106
by (typechk_tac (bin_typechecks0@integ_typechecks@
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   107
                 nat_typechecks@[bool_into_nat]));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   108
qed "integ_of_bin_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   109
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   110
Goalw [norm_Bcons_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   111
    "[| w: bin; b: bool |] ==> norm_Bcons(w,b) : bin";
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   112
by (etac bin.elim 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   113
by (ALLGOALS (asm_simp_tac (simpset() addsimps bin.case_eqns)));
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   114
by (typechk_tac (bin_typechecks0@bool_typechecks));
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   115
qed "norm_Bcons_type";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   116
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   117
Goalw [bin_succ_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   118
    "w: bin ==> bin_succ(w) : bin";
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   119
by (typechk_tac ([norm_Bcons_type]@bin_typechecks0@bool_typechecks));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   120
qed "bin_succ_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   121
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   122
Goalw [bin_pred_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   123
    "w: bin ==> bin_pred(w) : bin";
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   124
by (typechk_tac ([norm_Bcons_type]@bin_typechecks0@bool_typechecks));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   125
qed "bin_pred_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   126
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   127
Goalw [bin_minus_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   128
    "w: bin ==> bin_minus(w) : bin";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   129
by (typechk_tac ([bin_pred_type]@bin_typechecks0@bool_typechecks));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   130
qed "bin_minus_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   131
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   132
Goalw [bin_add_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   133
    "[| 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
   134
by (typechk_tac ([norm_Bcons_type, bin_succ_type, bin_pred_type]@
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   135
                 bin_typechecks0@ bool_typechecks@ZF_typechecks));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   136
qed "bin_add_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   137
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   138
Goalw [bin_mult_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   139
    "[| 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
   140
by (typechk_tac ([norm_Bcons_type, bin_minus_type, bin_add_type]@
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   141
                 bin_typechecks0@ bool_typechecks));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   142
qed "bin_mult_type";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   143
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   144
val bin_typechecks = bin_typechecks0 @
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   145
    [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
   146
     bin_minus_type, bin_add_type, bin_mult_type];
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   147
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   148
Addsimps ([bool_1I, bool_0I,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   149
	   bin_rec_Plus, bin_rec_Minus, bin_rec_Bcons] @ 
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   150
	  bin_recs integ_of_bin_def @ bin_typechecks);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   151
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   152
val typechecks = bin_typechecks @ integ_typechecks @ nat_typechecks @
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   153
                 [bool_subset_nat RS subsetD];
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
(**** The carry/borrow functions, bin_succ and bin_pred ****)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   156
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   157
(** Lemmas **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   158
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   159
goal Integ.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   160
    "!!z v. [| z $+ v = z' $+ v';  \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   161
\       z: integ; z': integ;  v: integ; v': integ;  w: integ |]   \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   162
\    ==> z $+ (v $+ w) = z' $+ (v' $+ w)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   163
by (asm_simp_tac (simpset() addsimps ([zadd_assoc RS sym])) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   164
qed "zadd_assoc_cong";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   165
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   166
goal Integ.thy 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   167
    "!!z v w. [| z: integ;  v: integ;  w: integ |]   \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   168
\    ==> z $+ (v $+ w) = v $+ (z $+ w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   169
by (REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   170
qed "zadd_assoc_swap";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   171
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   172
(*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
   173
bind_thm ("zadd_assoc_znat", (znat_type RS zadd_assoc_swap));
515
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
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   176
Addsimps (bin_recs bin_succ_def @ bin_recs bin_pred_def);
515
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*)
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   180
Goal "[| w: bin; b: bool |] ==>     \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   181
\         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
   182
by (etac bin.elim 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   183
by (asm_simp_tac (simpset() addsimps norm_Bcons_simps) 3);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   184
by (ALLGOALS (etac boolE));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   185
by (ALLGOALS (asm_simp_tac (simpset() addsimps (norm_Bcons_simps))));
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   186
qed "integ_of_bin_norm_Bcons";
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   187
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   188
Goal "w: bin ==> integ_of_bin(bin_succ(w)) = $#1 $+ integ_of_bin(w)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   189
by (etac bin.induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   190
by (Simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   191
by (Simp_tac 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   192
by (etac boolE 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   193
by (ALLGOALS
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   194
    (asm_simp_tac (simpset() addsimps integ_of_bin_norm_Bcons::zadd_ac)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   195
qed "integ_of_bin_succ";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   196
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   197
Goal "w: bin ==> integ_of_bin(bin_pred(w)) = $~ ($#1) $+ integ_of_bin(w)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   198
by (etac bin.induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   199
by (Simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   200
by (Simp_tac 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   201
by (etac boolE 1);
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   202
by (ALLGOALS
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   203
    (asm_simp_tac (simpset() addsimps integ_of_bin_norm_Bcons::zadd_ac)));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   204
qed "integ_of_bin_pred";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   205
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   206
(*These two results replace the definitions of bin_succ and bin_pred*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   207
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
(*** bin_minus: (unary!) negation of binary integers ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   210
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   211
Addsimps (bin_recs bin_minus_def @
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   212
	  [integ_of_bin_succ, integ_of_bin_pred]);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   213
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   214
Goal "w: bin ==> integ_of_bin(bin_minus(w)) = $~ integ_of_bin(w)";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   215
by (etac bin.induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   216
by (Simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   217
by (Simp_tac 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   218
by (etac boolE 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   219
by (ALLGOALS 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   220
    (asm_simp_tac (simpset() addsimps (zadd_ac@[zminus_zadd_distrib]))));
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   221
qed "integ_of_bin_minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   222
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   223
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   224
(*** bin_add: binary addition ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   225
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   226
Goalw [bin_add_def] "w: bin ==> bin_add(Plus,w) = w";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   227
by (Asm_simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   228
qed "bin_add_Plus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   229
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5068
diff changeset
   230
Goalw [bin_add_def] "w: bin ==> bin_add(Minus,w) = bin_pred(w)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   231
by (Asm_simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   232
qed "bin_add_Minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   233
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   234
Goalw [bin_add_def] "bin_add(Bcons(v,x),Plus) = Bcons(v,x)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   235
by (Simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   236
qed "bin_add_Bcons_Plus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   237
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   238
Goalw [bin_add_def] "bin_add(Bcons(v,x),Minus) = bin_pred(Bcons(v,x))";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   239
by (Simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 515
diff changeset
   240
qed "bin_add_Bcons_Minus";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   241
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   242
Goalw [bin_add_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   243
    "[| w: bin;  y: bool |] ==> \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   244
\           bin_add(Bcons(v,x), Bcons(w,y)) = \
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   245
\           norm_Bcons(bin_add(v, cond(x and y, bin_succ(w), w)), x xor y)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   246
by (Asm_simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   247
qed "bin_add_Bcons_Bcons";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   248
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   249
Addsimps [bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   250
	  bin_add_Bcons_Minus, bin_add_Bcons_Bcons,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   251
	  integ_of_bin_succ, integ_of_bin_pred,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   252
	  integ_of_bin_norm_Bcons];
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   253
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   254
Addsimps [bool_subset_nat RS subsetD];
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   255
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   256
Goal "v: bin ==> \
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   257
\         ALL w: bin. integ_of_bin(bin_add(v,w)) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   258
\                     integ_of_bin(v) $+ integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   259
by (etac bin.induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   260
by (Simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   261
by (Simp_tac 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   262
by (rtac ballI 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   263
by (bin_ind_tac "wa" [] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   264
by (ALLGOALS (asm_simp_tac (simpset() addsimps zadd_ac setloop (etac boolE))));
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   265
val integ_of_bin_add_lemma = result();
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   266
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   267
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
   268
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   269
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   270
(*** bin_add: binary multiplication ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   271
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   272
Addsimps (bin_recs bin_mult_def @ 
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   273
	  [integ_of_bin_minus, integ_of_bin_add,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   274
	   integ_of_bin_norm_Bcons]);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   275
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   276
val major::prems = goal Bin.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   277
    "[| v: bin; w: bin |] ==>   \
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   278
\    integ_of_bin(bin_mult(v,w)) = \
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   279
\    integ_of_bin(v) $* integ_of_bin(w)";
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   280
by (cut_facts_tac prems 1);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   281
by (bin_ind_tac "v" [major] 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   282
by (Asm_simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   283
by (Asm_simp_tac 1);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   284
by (etac boolE 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   285
by (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib]) 2);
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   286
by (asm_simp_tac 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   287
    (simpset() addsimps ([zadd_zmult_distrib, zmult_1] @ zadd_ac)) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   288
qed "integ_of_bin_mult";
515
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
(**** Computations ****)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   291
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   292
(** extra rules for bin_succ, bin_pred **)
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
val [bin_succ_Plus, bin_succ_Minus, _] = bin_recs bin_succ_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   295
val [bin_pred_Plus, bin_pred_Minus, _] = bin_recs bin_pred_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   296
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   297
Goal "bin_succ(Bcons(w,1)) = Bcons(bin_succ(w), 0)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   298
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   299
qed "bin_succ_Bcons1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   300
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   301
Goal "bin_succ(Bcons(w,0)) = norm_Bcons(w,1)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   302
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   303
qed "bin_succ_Bcons0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   304
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   305
Goal "bin_pred(Bcons(w,1)) = norm_Bcons(w,0)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   306
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   307
qed "bin_pred_Bcons1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   308
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   309
Goal "bin_pred(Bcons(w,0)) = Bcons(bin_pred(w), 1)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   310
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   311
qed "bin_pred_Bcons0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   312
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   313
(** extra rules for bin_minus **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   314
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   315
val [bin_minus_Plus, bin_minus_Minus, _] = bin_recs bin_minus_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   316
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   317
Goal "bin_minus(Bcons(w,1)) = bin_pred(Bcons(bin_minus(w), 0))";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   318
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   319
qed "bin_minus_Bcons1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   320
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   321
Goal "bin_minus(Bcons(w,0)) = Bcons(bin_minus(w), 0)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   322
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   323
qed "bin_minus_Bcons0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   324
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   325
(** extra rules for bin_add **)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   326
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   327
Goal "w: bin ==> bin_add(Bcons(v,1), Bcons(w,1)) = \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   328
\                    norm_Bcons(bin_add(v, bin_succ(w)), 0)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   329
by (Asm_simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   330
qed "bin_add_Bcons_Bcons11";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   331
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   332
Goal "w: bin ==> bin_add(Bcons(v,1), Bcons(w,0)) =  \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   333
\                    norm_Bcons(bin_add(v,w), 1)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   334
by (Asm_simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   335
qed "bin_add_Bcons_Bcons10";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   336
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   337
Goal "[| w: bin;  y: bool |] ==> \
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   338
\           bin_add(Bcons(v,0), Bcons(w,y)) = norm_Bcons(bin_add(v,w), y)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   339
by (Asm_simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   340
qed "bin_add_Bcons_Bcons0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   341
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   342
(** extra rules for bin_mult **)
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
val [bin_mult_Plus, bin_mult_Minus, _] = bin_recs bin_mult_def;
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   345
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   346
Goal "bin_mult(Bcons(v,1), w) = bin_add(norm_Bcons(bin_mult(v,w),0), w)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   347
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   348
qed "bin_mult_Bcons1";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   349
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   350
Goal "bin_mult(Bcons(v,0), w) = norm_Bcons(bin_mult(v,w),0)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   351
by (Simp_tac 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   352
qed "bin_mult_Bcons0";
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   353
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   354
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   355
(*** The computation simpset ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   356
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 2637
diff changeset
   357
val bin_comp_ss = simpset_of Integ.thy 
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   358
    addsimps [integ_of_bin_add RS sym,   (*invoke bin_add*)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   359
              integ_of_bin_minus RS sym, (*invoke bin_minus*)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   360
              integ_of_bin_mult RS sym,  (*invoke bin_mult*)
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   361
              bin_succ_Plus, bin_succ_Minus,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   362
              bin_succ_Bcons1, bin_succ_Bcons0,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   363
              bin_pred_Plus, bin_pred_Minus,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   364
              bin_pred_Bcons1, bin_pred_Bcons0,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   365
              bin_minus_Plus, bin_minus_Minus,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   366
              bin_minus_Bcons1, bin_minus_Bcons0,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   367
              bin_add_Plus, bin_add_Minus, bin_add_Bcons_Plus, 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   368
              bin_add_Bcons_Minus, bin_add_Bcons_Bcons0, 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   369
              bin_add_Bcons_Bcons10, bin_add_Bcons_Bcons11,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   370
              bin_mult_Plus, bin_mult_Minus,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   371
              bin_mult_Bcons1, bin_mult_Bcons0] @
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   372
             norm_Bcons_simps
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2469
diff changeset
   373
    setSolver (type_auto_tac ([bool_1I, bool_0I] @ bin_typechecks0));
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   374
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
   375
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   376
(*** Examples of performing binary arithmetic by simplification ***)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   377
4446
wenzelm
parents: 4091
diff changeset
   378
set proof_timing;
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   379
(*All runtimes below are on a SPARCserver 10*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   380
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   381
Goal "#13  $+  #19 = #32";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   382
by (simp_tac bin_comp_ss 1);    (*0.4 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   383
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   384
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   385
bin_add(binary_of_int 13, binary_of_int 19);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   386
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   387
Goal "#1234  $+  #5678 = #6912";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   388
by (simp_tac bin_comp_ss 1);    (*1.3 secs*)
515
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 1234, binary_of_int 5678);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   392
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   393
Goal "#1359  $+  #~2468 = #~1109";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   394
by (simp_tac bin_comp_ss 1);    (*1.2 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   395
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   396
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   397
bin_add(binary_of_int 1359, binary_of_int ~2468);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   398
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   399
Goal "#93746  $+  #~46375 = #47371";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   400
by (simp_tac bin_comp_ss 1);    (*1.9 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   401
result();
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
bin_add(binary_of_int 93746, binary_of_int ~46375);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   404
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   405
Goal "$~ #65745 = #~65745";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   406
by (simp_tac bin_comp_ss 1);    (*0.4 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   407
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   408
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   409
bin_minus(binary_of_int 65745);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   410
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   411
(* negation of ~54321 *)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   412
Goal "$~ #~54321 = #54321";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   413
by (simp_tac bin_comp_ss 1);    (*0.5 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   414
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   415
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   416
bin_minus(binary_of_int ~54321);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   417
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   418
Goal "#13  $*  #19 = #247";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   419
by (simp_tac bin_comp_ss 1);    (*0.7 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   420
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   421
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   422
bin_mult(binary_of_int 13, binary_of_int 19);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   423
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   424
Goal "#~84  $*  #51 = #~4284";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   425
by (simp_tac bin_comp_ss 1);    (*1.3 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   426
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   427
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   428
bin_mult(binary_of_int ~84, binary_of_int 51);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   429
906
6cd9c397f36a Redefined arithmetic to suppress redundant leading 0s and 1s
lcp
parents: 782
diff changeset
   430
(*The worst case for 8-bit operands *)
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   431
Goal "#255  $*  #255 = #65025";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   432
by (simp_tac bin_comp_ss 1);    (*4.3 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   433
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   434
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   435
bin_mult(binary_of_int 255, binary_of_int 255);
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   436
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4446
diff changeset
   437
Goal "#1359  $*  #~2468 = #~3354012";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 906
diff changeset
   438
by (simp_tac bin_comp_ss 1);    (*6.1 secs*)
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   439
result();
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   440
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents: 477
diff changeset
   441
bin_mult(binary_of_int 1359, binary_of_int ~2468);