src/ZF/Arith.ML
author wenzelm
Thu, 15 Nov 2001 18:20:13 +0100
changeset 12207 4dff931b852f
parent 11386 cf8d81cf8034
child 12789 459b5de466b2
permissions -rw-r--r--
added Induct/Binary_Trees.thy, Induct/Tree_Forest (converted from former ex/TF.ML ex/TF.thy ex/Term.ML ex/Term.thy);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
     1
(*  Title:      ZF/Arith.ML
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
     6
Arithmetic operators and their definitions
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
Proofs about elementary arithmetic: addition, multiplication, etc.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
(*"Difference" is subtraction of natural numbers.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
  There are no negative numbers; we have
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
     m #- n = 0  iff  m<=n   and     m #- n = succ(k) iff m>n.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
  Also, rec(m, 0, %z w.z) is pred(m).   
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
    17
Addsimps [rec_type, nat_0_le];
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9548
diff changeset
    18
bind_thms ("nat_typechecks", [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    20
Goal "[| 0<k; k: nat |] ==> EX j: nat. k = succ(j)";
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
    21
by (etac rev_mp 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    22
by (induct_tac "k" 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    23
by (Simp_tac 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
    24
by (Blast_tac 1);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
    25
val lemma = result();
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
    26
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
    27
(* [| 0 < k; k: nat; !!j. [| j: nat; k = succ(j) |] ==> Q |] ==> Q *)
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
    28
bind_thm ("zero_lt_natE", lemma RS bexE);
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
    29
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
    30
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    31
(** natify: coercion to "nat" **)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    32
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    33
Goalw [pred_def] "pred(succ(y)) = y";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    34
by Auto_tac;  
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    35
qed "pred_succ_eq";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    36
Addsimps [pred_succ_eq];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    37
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    38
Goal "natify(succ(x)) = succ(natify(x))";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    39
by (rtac (natify_def RS def_Vrecursor RS trans) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    40
by Auto_tac;  
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    41
qed "natify_succ";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    42
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    43
Goal "natify(0) = 0";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    44
by (rtac (natify_def RS def_Vrecursor RS trans) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    45
by Auto_tac;  
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    46
qed "natify_0";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    47
Addsimps [natify_0];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    48
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    49
Goal "ALL z. x ~= succ(z) ==> natify(x) = 0";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    50
by (rtac (natify_def RS def_Vrecursor RS trans) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    51
by Auto_tac;  
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    52
qed "natify_non_succ";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    53
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    54
Goal "natify(x) : nat";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    55
by (eps_ind_tac "x" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    56
by (case_tac "EX z. x1 = succ(z)" 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
    57
by (auto_tac (claset(), simpset() addsimps [natify_succ, natify_non_succ]));  
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    58
qed "natify_in_nat";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    59
AddIffs [natify_in_nat];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    60
AddTCs [natify_in_nat];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    61
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    62
Goal "n : nat ==> natify(n) = n";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    63
by (induct_tac "n" 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
    64
by (auto_tac (claset(), simpset() addsimps [natify_succ]));  
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    65
qed "natify_ident";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    66
Addsimps [natify_ident];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    67
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
    68
Goal "[|natify(x) = y;  x: nat|] ==> x=y";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
    69
by Auto_tac; 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
    70
qed "natify_eqE";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
    71
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    72
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    73
(*** Collapsing rules: to remove natify from arithmetic expressions ***)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    74
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    75
Goal "natify(natify(x)) = natify(x)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    76
by (Simp_tac 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    77
qed "natify_idem";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    78
Addsimps [natify_idem];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    79
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
(** Addition **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    82
Goal "natify(m) #+ n = m #+ n";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    83
by (simp_tac (simpset() addsimps [add_def]) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    84
qed "add_natify1";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    85
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    86
Goal "m #+ natify(n) = m #+ n";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    87
by (simp_tac (simpset() addsimps [add_def]) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    88
qed "add_natify2";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    89
Addsimps [add_natify1, add_natify2];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    90
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    91
(** Multiplication **)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    92
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    93
Goal "natify(m) #* n = m #* n";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    94
by (simp_tac (simpset() addsimps [mult_def]) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    95
qed "mult_natify1";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    96
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    97
Goal "m #* natify(n) = m #* n";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    98
by (simp_tac (simpset() addsimps [mult_def]) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
    99
qed "mult_natify2";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   100
Addsimps [mult_natify1, mult_natify2];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   101
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   102
(** Difference **)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   103
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   104
Goal "natify(m) #- n = m #- n";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   105
by (simp_tac (simpset() addsimps [diff_def]) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   106
qed "diff_natify1";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   107
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   108
Goal "m #- natify(n) = m #- n";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   109
by (simp_tac (simpset() addsimps [diff_def]) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   110
qed "diff_natify2";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   111
Addsimps [diff_natify1, diff_natify2];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   112
9492
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   113
(** Remainder **)
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   114
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   115
Goal "natify(m) mod n = m mod n";
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   116
by (simp_tac (simpset() addsimps [mod_def]) 1);
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   117
qed "mod_natify1";
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   118
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   119
Goal "m mod natify(n) = m mod n";
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   120
by (simp_tac (simpset() addsimps [mod_def]) 1);
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   121
qed "mod_natify2";
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   122
Addsimps [mod_natify1, mod_natify2];
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   123
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   124
(** Quotient **)
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   125
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   126
Goal "natify(m) div n = m div n";
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   127
by (simp_tac (simpset() addsimps [div_def]) 1);
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   128
qed "div_natify1";
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   129
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   130
Goal "m div natify(n) = m div n";
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   131
by (simp_tac (simpset() addsimps [div_def]) 1);
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   132
qed "div_natify2";
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   133
Addsimps [div_natify1, div_natify2];
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   134
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   135
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   136
(*** Typing rules ***)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   137
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   138
(** Addition **)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   139
9492
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   140
Goal "[| m:nat;  n:nat |] ==> raw_add (m, n) : nat";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   141
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   142
by Auto_tac;
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   143
qed "raw_add_type";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   144
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   145
Goal "m #+ n : nat";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   146
by (simp_tac (simpset() addsimps [add_def, raw_add_type]) 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   147
qed "add_type";
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   148
AddIffs [add_type];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
   149
AddTCs [add_type];
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   150
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   151
(** Multiplication **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   152
9492
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   153
Goal "[| m:nat;  n:nat |] ==> raw_mult (m, n) : nat";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   154
by (induct_tac "m" 1);
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   155
by (ALLGOALS (asm_simp_tac (simpset() addsimps [raw_add_type])));
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   156
qed "raw_mult_type";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   157
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   158
Goal "m #* n : nat";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   159
by (simp_tac (simpset() addsimps [mult_def, raw_mult_type]) 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   160
qed "mult_type";
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   161
AddIffs [mult_type];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
   162
AddTCs [mult_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   164
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   165
(** Difference **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   166
9492
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   167
Goal "[| m:nat;  n:nat |] ==> raw_diff (m, n) : nat";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   168
by (induct_tac "n" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   169
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   170
by (fast_tac (claset() addIs [nat_case_type]) 1);
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   171
qed "raw_diff_type";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   172
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   173
Goal "m #- n : nat";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   174
by (simp_tac (simpset() addsimps [diff_def, raw_diff_type]) 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   175
qed "diff_type";
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   176
AddIffs [diff_type];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
   177
AddTCs [diff_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   178
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   179
Goalw [diff_def] "0 #- n = 0";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   180
by (rtac (natify_in_nat RS nat_induct) 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   181
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   182
qed "diff_0_eq_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   183
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   184
(*Must simplify BEFORE the induction: else we get a critical pair*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   185
Goal "succ(m) #- succ(n) = m #- n";
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   186
by (simp_tac (simpset() addsimps [natify_succ, diff_def]) 1);
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   187
by (res_inst_tac [("x1","n")] (natify_in_nat RS nat_induct) 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   188
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   189
qed "diff_succ_succ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   190
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   191
(*This defining property is no longer wanted*)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   192
Delsimps [raw_diff_succ];  
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   193
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   194
(*Natify has weakened this law, compared with the older approach*)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   195
Goal "m #- 0 = natify(m)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   196
by (asm_simp_tac (simpset() addsimps [diff_def]) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   197
qed "diff_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   198
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   199
Addsimps [diff_0, diff_0_eq_0, diff_succ_succ];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   200
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   201
Goal "m:nat ==> (m #- n) le m";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   202
by (subgoal_tac "(m #- natify(n)) le m" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   203
by (res_inst_tac [("m","m"), ("n","natify(n)")] diff_induct 2);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   204
by (etac leE 6);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   205
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [le_iff])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   206
qed "diff_le_self";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   207
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   209
(*** Addition ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   211
(*Natify has weakened this law, compared with the older approach*)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   212
Goal "0 #+ m = natify(m)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   213
by (asm_simp_tac (simpset() addsimps [add_def]) 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   214
qed "add_0_natify";
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   215
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   216
Goal "succ(m) #+ n = succ(m #+ n)";
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   217
by (asm_simp_tac (simpset() addsimps [natify_succ, add_def]) 1);
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   218
qed "add_succ";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   219
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   220
Addsimps [add_0_natify, add_succ];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   221
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   222
Goal "m: nat ==> 0 #+ m = m";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   223
by (Asm_simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   224
qed "add_0";
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   225
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   226
(*Associative law for addition*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   227
Goal "(m #+ n) #+ k = m #+ (n #+ k)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   228
by (subgoal_tac "(natify(m) #+ natify(n)) #+ natify(k) = \
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   229
\                natify(m) #+ (natify(n) #+ natify(k))" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   230
by (res_inst_tac [("n","natify(m)")] nat_induct 2);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   231
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   232
qed "add_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   233
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   234
(*The following two lemmas are used for add_commute and sometimes
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   235
  elsewhere, since they are safe for rewriting.*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   236
Goal "m #+ 0 = natify(m)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   237
by (subgoal_tac "natify(m) #+ 0 = natify(m)" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   238
by (res_inst_tac [("n","natify(m)")] nat_induct 2);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   239
by Auto_tac;
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   240
qed "add_0_right_natify";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   241
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   242
Goalw [add_def] "m #+ succ(n) = succ(m #+ n)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   243
by (res_inst_tac [("n","natify(m)")] nat_induct 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   244
by (auto_tac (claset(), simpset() addsimps [natify_succ]));  
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   245
qed "add_succ_right";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   246
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   247
Addsimps [add_0_right_natify, add_succ_right];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   248
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   249
Goal "m: nat ==> m #+ 0 = m";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   250
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   251
qed "add_0_right";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   252
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   253
(*Commutative law for addition*)  
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   254
Goal "m #+ n = n #+ m";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   255
by (subgoal_tac "natify(m) #+ natify(n) = natify(n) #+ natify(m)" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   256
by (res_inst_tac [("n","natify(m)")] nat_induct 2);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   257
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   258
qed "add_commute";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   259
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   260
(*for a/c rewriting*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   261
Goal "m#+(n#+k)=n#+(m#+k)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   262
by (rtac (add_commute RS trans) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   263
by (rtac (add_assoc RS trans) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   264
by (rtac (add_commute RS subst_context) 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   265
qed "add_left_commute";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   266
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   267
(*Addition is an AC-operator*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9548
diff changeset
   268
bind_thms ("add_ac", [add_assoc, add_commute, add_left_commute]);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   269
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
(*Cancellation law on the left*)
9492
72e429c66608 used natify with div and mod; also put in the divide-by-zero trick
paulson
parents: 9491
diff changeset
   271
Goal "[| raw_add(k, m) = raw_add(k, n);  k:nat |] ==> m=n";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   272
by (etac rev_mp 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   273
by (induct_tac "k" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   274
by Auto_tac;
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   275
qed "raw_add_left_cancel";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   276
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   277
Goalw [add_def] "k #+ m = k #+ n ==> natify(m) = natify(n)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   278
by (dtac raw_add_left_cancel 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   279
by Auto_tac;
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   280
qed "add_left_cancel_natify";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   281
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   282
Goal "[| i = j;  i #+ m = j #+ n;  m:nat;  n:nat |] ==> m = n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   283
by (force_tac (claset() addSDs [add_left_cancel_natify], simpset()) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   284
qed "add_left_cancel";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   286
(*Thanks to Sten Agerholm*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   287
Goal "k#+m le k#+n ==> natify(m) le natify(n)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   288
by (res_inst_tac [("P", "natify(k)#+m le natify(k)#+n")] rev_mp 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   289
by (res_inst_tac [("n","natify(k)")] nat_induct 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   290
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   291
qed "add_le_elim1_natify";
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   292
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   293
Goal "[| k#+m le k#+n; m: nat; n: nat |] ==> m le n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   294
by (dtac add_le_elim1_natify 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   295
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   296
qed "add_le_elim1";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   297
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   298
Goal "k#+m < k#+n ==> natify(m) < natify(n)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   299
by (res_inst_tac [("P", "natify(k)#+m < natify(k)#+n")] rev_mp 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   300
by (res_inst_tac [("n","natify(k)")] nat_induct 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   301
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   302
qed "add_lt_elim1_natify";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   303
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   304
Goal "[| k#+m < k#+n; m: nat; n: nat |] ==> m < n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   305
by (dtac add_lt_elim1_natify 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   306
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   307
qed "add_lt_elim1";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   308
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   309
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   310
(*** Monotonicity of Addition ***)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   311
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   312
(*strict, in 1st argument; proof is by rule induction on 'less than'.
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   313
  Still need j:nat, for consider j = omega.  Then we can have i<omega,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   314
  which is the same as i:nat, but natify(j)=0, so the conclusion fails.*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   315
Goal "[| i<j; j:nat |] ==> i#+k < j#+k";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   316
by (ftac lt_nat_in_nat 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   317
by (assume_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   318
by (etac succ_lt_induct 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   319
by (ALLGOALS (asm_simp_tac (simpset() addsimps [leI])));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   320
qed "add_lt_mono1";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   321
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   322
(*strict, in both arguments*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   323
Goal "[| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   324
by (rtac (add_lt_mono1 RS lt_trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   325
by (REPEAT (assume_tac 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   326
by (EVERY [stac add_commute 1,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   327
           stac add_commute 1,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   328
           rtac add_lt_mono1 1]);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   329
by (REPEAT (assume_tac 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   330
qed "add_lt_mono";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   331
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   332
(*A [clumsy] way of lifting < monotonicity to le monotonicity *)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   333
val lt_mono::ford::prems = Goal
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   334
     "[| !!i j. [| i<j; j:k |] ==> f(i) < f(j); \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   335
\        !!i. i:k ==> Ord(f(i));                \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   336
\        i le j;  j:k                           \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   337
\     |] ==> f(i) le f(j)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   338
by (cut_facts_tac prems 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   339
by (blast_tac (le_cs addSIs [lt_mono,ford] addSEs [leE]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   340
qed "Ord_lt_mono_imp_le_mono";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   341
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   342
(*le monotonicity, 1st argument*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   343
Goal "[| i le j; j:nat |] ==> i#+k le j#+k";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   344
by (res_inst_tac [("f", "%j. j#+k")] Ord_lt_mono_imp_le_mono 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   345
by (REPEAT (ares_tac [add_lt_mono1, add_type RS nat_into_Ord] 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   346
qed "add_le_mono1";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   347
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   348
(* le monotonicity, BOTH arguments*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   349
Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   350
by (rtac (add_le_mono1 RS le_trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   351
by (REPEAT (assume_tac 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   352
by (EVERY [stac add_commute 1,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   353
           stac add_commute 1,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   354
           rtac add_le_mono1 1]);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   355
by (REPEAT (assume_tac 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   356
qed "add_le_mono";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   357
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   358
(** Subtraction is the inverse of addition. **)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   359
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   360
Goal "(n#+m) #- n = natify(m)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   361
by (subgoal_tac "(natify(n) #+ m) #- natify(n) = natify(m)" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   362
by (res_inst_tac [("n","natify(n)")] nat_induct 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   363
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   364
qed "diff_add_inverse";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   365
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   366
Goal "(m#+n) #- n = natify(m)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   367
by (simp_tac (simpset() addsimps [inst "m" "m" add_commute, 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   368
				  diff_add_inverse]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   369
qed "diff_add_inverse2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   370
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   371
Goal "(k#+m) #- (k#+n) = m #- n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   372
by (subgoal_tac "(natify(k) #+ natify(m)) #- (natify(k) #+ natify(n)) = \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   373
\                natify(m) #- natify(n)" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   374
by (res_inst_tac [("n","natify(k)")] nat_induct 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   375
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   376
qed "diff_cancel";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   377
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   378
Goal "(m#+k) #- (n#+k) = m #- n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   379
by (simp_tac (simpset() addsimps [inst "n" "k" add_commute, diff_cancel]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   380
qed "diff_cancel2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   381
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   382
Goal "n #- (n#+m) = 0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   383
by (subgoal_tac "natify(n) #- (natify(n) #+ natify(m)) = 0" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   384
by (res_inst_tac [("n","natify(n)")] nat_induct 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   385
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   386
qed "diff_add_0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   387
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   388
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   389
(** Lemmas for the CancelNumerals simproc **)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   390
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   391
Goal "(u #+ m = u #+ n) <-> (0 #+ m = natify(n))";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   392
by Auto_tac;  
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   393
by (blast_tac (claset() addDs [add_left_cancel_natify]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   394
by (asm_full_simp_tac (simpset() addsimps [add_def]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   395
qed "eq_add_iff";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   396
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   397
Goal "(u #+ m < u #+ n) <-> (0 #+ m < natify(n))";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   398
by (auto_tac (claset(), simpset() addsimps [add_lt_elim1_natify]));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   399
by (dtac add_lt_mono1 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   400
by (auto_tac (claset(), simpset() addsimps [inst "m" "u" add_commute]));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   401
qed "less_add_iff";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   402
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   403
Goal "((u #+ m) #- (u #+ n)) = ((0 #+ m) #- n)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   404
by (asm_simp_tac (simpset() addsimps [diff_cancel]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   405
qed "diff_add_eq";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   406
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   407
(*To tidy up the result of a simproc.  Only the RHS will be simplified.*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   408
Goal "u = u' ==> (t==u) == (t==u')";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   409
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   410
qed "eq_cong2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   411
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   412
Goal "u <-> u' ==> (t==u) == (t==u')";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   413
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   414
qed "iff_cong2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   415
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   416
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   417
(*** Multiplication [the simprocs need these laws] ***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   418
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   419
Goal "0 #* m = 0";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   420
by (simp_tac (simpset() addsimps [mult_def]) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   421
qed "mult_0";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   422
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   423
Goal "succ(m) #* n = n #+ (m #* n)";
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   424
by (simp_tac (simpset() addsimps [add_def, mult_def, natify_succ, 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   425
                                  raw_mult_type]) 1);
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   426
qed "mult_succ";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   427
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   428
Addsimps [mult_0, mult_succ];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   429
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   430
(*right annihilation in product*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   431
Goalw [mult_def] "m #* 0 = 0";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   432
by (res_inst_tac [("n","natify(m)")] nat_induct 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   433
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   434
qed "mult_0_right";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   435
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   436
(*right successor law for multiplication*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   437
Goal "m #* succ(n) = m #+ (m #* n)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   438
by (subgoal_tac "natify(m) #* succ(natify(n)) = \
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   439
\                natify(m) #+ (natify(m) #* natify(n))" 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9492
diff changeset
   440
by (full_simp_tac (simpset() addsimps [natify_succ, add_def, mult_def]) 1);
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   441
by (res_inst_tac [("n","natify(m)")] nat_induct 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   442
by (ALLGOALS (asm_simp_tac (simpset() addsimps add_ac)));
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   443
qed "mult_succ_right";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   444
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   445
Addsimps [mult_0_right, mult_succ_right];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   446
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   447
Goal "1 #* n = natify(n)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   448
by Auto_tac;
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   449
qed "mult_1_natify";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   450
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   451
Goal "n #* 1 = natify(n)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   452
by Auto_tac;
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   453
qed "mult_1_right_natify";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   454
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   455
Addsimps [mult_1_natify, mult_1_right_natify];
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   456
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   457
Goal "n : nat ==> 1 #* n = n";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   458
by (Asm_simp_tac 1);
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   459
qed "mult_1";
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   460
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   461
Goal "n : nat ==> n #* 1 = n";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   462
by (Asm_simp_tac 1);
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   463
qed "mult_1_right";
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   464
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   465
(*Commutative law for multiplication*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   466
Goal "m #* n = n #* m";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   467
by (subgoal_tac "natify(m) #* natify(n) = natify(n) #* natify(m)" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   468
by (res_inst_tac [("n","natify(m)")] nat_induct 2);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   469
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   470
qed "mult_commute";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   471
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   472
(*addition distributes over multiplication*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   473
Goal "(m #+ n) #* k = (m #* k) #+ (n #* k)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   474
by (subgoal_tac "(natify(m) #+ natify(n)) #* natify(k) = \
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   475
\                (natify(m) #* natify(k)) #+ (natify(n) #* natify(k))" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   476
by (res_inst_tac [("n","natify(m)")] nat_induct 2);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   477
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_assoc RS sym])));
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   478
qed "add_mult_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   479
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   480
(*Distributive law on the left*)
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   481
Goal "k #* (m #+ n) = (k #* m) #+ (k #* n)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   482
by (subgoal_tac "natify(k) #* (natify(m) #+ natify(n)) = \
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   483
\                (natify(k) #* natify(m)) #+ (natify(k) #* natify(n))" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   484
by (res_inst_tac [("n","natify(m)")] nat_induct 2);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   485
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps add_ac)));
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   486
qed "add_mult_distrib_left";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   487
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   488
(*Associative law for multiplication*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   489
Goal "(m #* n) #* k = m #* (n #* k)";
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   490
by (subgoal_tac "(natify(m) #* natify(n)) #* natify(k) = \
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   491
\                natify(m) #* (natify(n) #* natify(k))" 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   492
by (res_inst_tac [("n","natify(m)")] nat_induct 2);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   493
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_mult_distrib])));
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   494
qed "mult_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   495
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   496
(*for a/c rewriting*)
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   497
Goal "m #* (n #* k) = n #* (m #* k)";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   498
by (rtac (mult_commute RS trans) 1);
9491
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   499
by (rtac (mult_assoc RS trans) 1);
1a36151ee2fc natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents: 9301
diff changeset
   500
by (rtac (mult_commute RS subst_context) 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   501
qed "mult_left_commute";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   502
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9548
diff changeset
   503
bind_thms ("mult_ac", [mult_assoc,mult_commute,mult_left_commute]);