src/ZF/Arith.ML
author wenzelm
Sat, 01 Jul 2000 19:55:22 +0200
changeset 9230 17ae63f82ad8
parent 8551 5c22595bc599
child 9301 de04717eed78
permissions -rw-r--r--
GPLed;
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];
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
    18
val 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
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
(** Addition **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    33
Goal "[| m:nat;  n:nat |] ==> m #+ n : nat";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    34
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    35
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    36
qed "add_type";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    37
Addsimps [add_type];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
    38
AddTCs [add_type];
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    39
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
(** Multiplication **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    41
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    42
Goal "[| m:nat;  n:nat |] ==> m #* n : nat";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    43
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    44
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    45
qed "mult_type";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    46
Addsimps [mult_type];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
    47
AddTCs [mult_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    49
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
(** Difference **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    51
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    52
Goal "[| m:nat;  n:nat |] ==> m #- n : nat";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    53
by (induct_tac "n" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    54
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    55
by (fast_tac (claset() addIs [nat_case_type]) 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    56
qed "diff_type";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    57
Addsimps [diff_type];
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
    58
AddTCs [diff_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    60
Goal "n:nat ==> 0 #- n = 0";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    61
by (induct_tac "n" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    62
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    63
qed "diff_0_eq_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    65
(*Must simplify BEFORE the induction: else we get a critical pair*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    66
Goal "[| m:nat;  n:nat |] ==> succ(m) #- succ(n) = m #- n";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    67
by (Asm_simp_tac 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    68
by (induct_tac "n" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    69
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    70
qed "diff_succ_succ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    72
Addsimps [diff_0_eq_0, diff_succ_succ];
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    73
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    74
(*This defining property is no longer wanted*)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    75
Delsimps [diff_SUCC];  
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    76
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
val prems = goal Arith.thy 
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
    78
    "[| m:nat;  n:nat |] ==> m #- n le m";
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
    79
by (rtac (prems MRS diff_induct) 1);
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
    80
by (etac leE 3);
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
    81
by (ALLGOALS (asm_simp_tac (simpset() addsimps prems @ [le_iff])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
    82
qed "diff_le_self";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
(*** Simplification over add, mult, diff ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    85
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
val arith_typechecks = [add_type, mult_type, diff_type];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
(*** Addition ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
(*Associative law for addition*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    92
Goal "m:nat ==> (m #+ n) #+ k = m #+ (n #+ k)";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    93
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    94
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    95
qed "add_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    97
(*The following two lemmas are used for add_commute and sometimes
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
  elsewhere, since they are safe for rewriting.*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
    99
Goal "m:nat ==> m #+ 0 = m";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   100
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   101
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   102
qed "add_0_right";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   103
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   104
Goal "m:nat ==> m #+ succ(n) = succ(m #+ n)";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   105
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   106
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   107
qed "add_succ_right";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   108
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   109
Addsimps [add_0_right, add_succ_right];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   111
(*Commutative law for addition*)  
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   112
Goal "[| m:nat;  n:nat |] ==> m #+ n = n #+ m";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   113
by (induct_tac "n" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   114
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   115
qed "add_commute";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   116
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   117
(*for a/c rewriting*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   118
Goal "[| m:nat;  n:nat |] ==> m#+(n#+k)=n#+(m#+k)";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   119
by (asm_simp_tac (simpset() addsimps [add_assoc RS sym, add_commute]) 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   120
qed "add_left_commute";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   121
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   122
(*Addition is an AC-operator*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   123
val add_ac = [add_assoc, add_commute, add_left_commute];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   124
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   125
(*Cancellation law on the left*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   126
Goal "[| k #+ m = k #+ n;  k:nat |] ==> m=n";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   127
by (etac rev_mp 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   128
by (induct_tac "k" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   129
by Auto_tac;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   130
qed "add_left_cancel";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   131
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
(*** Multiplication ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
(*right annihilation in product*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   135
Goal "m:nat ==> m #* 0 = 0";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   136
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   137
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   138
qed "mult_0_right";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   140
(*right successor law for multiplication*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   141
Goal "[| m:nat;  n:nat |] ==> m #* succ(n) = m #+ (m #* n)";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   142
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   143
by (ALLGOALS (asm_simp_tac (simpset() addsimps add_ac)));
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   144
qed "mult_succ_right";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   145
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   146
Addsimps [mult_0_right, mult_succ_right];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   147
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   148
Goal "n:nat ==> 1 #* n = n";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   149
by (Asm_simp_tac 1);
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   150
qed "mult_1";
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   151
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   152
Goal "n:nat ==> n #* 1 = n";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   153
by (Asm_simp_tac 1);
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   154
qed "mult_1_right";
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   155
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   156
Addsimps [mult_1, mult_1_right];
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   157
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   158
(*Commutative law for multiplication*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   159
Goal "[| m:nat;  n:nat |] ==> m #* n = n #* m";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   160
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   161
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   162
qed "mult_commute";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   164
(*addition distributes over multiplication*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   165
Goal "[| m:nat;  k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   166
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   167
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_assoc RS sym])));
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   168
qed "add_mult_distrib";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   169
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   170
(*Distributive law on the left; requires an extra typing premise*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   171
Goal "[| m:nat;  n:nat;  k:nat |] ==> k #* (m #+ n) = (k #* m) #+ (k #* n)";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   172
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   173
by (ALLGOALS (asm_simp_tac (simpset() addsimps add_ac)));
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   174
qed "add_mult_distrib_left";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   175
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   176
(*Associative law for multiplication*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   177
Goal "[| m:nat;  n:nat;  k:nat |] ==> (m #* n) #* k = m #* (n #* k)";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   178
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   179
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_mult_distrib])));
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   180
qed "mult_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   181
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   182
(*for a/c rewriting*)
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   183
Goal "[| m:nat;  n:nat;  k:nat |] ==> m #* (n #* k) = n #* (m #* k)";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   184
by (rtac (mult_commute RS trans) 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   185
by (rtac (mult_assoc RS trans) 3);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   186
by (rtac (mult_commute RS subst_context) 6);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   187
by (REPEAT (ares_tac [mult_type] 1));
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   188
qed "mult_left_commute";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   189
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   190
val mult_ac = [mult_assoc,mult_commute,mult_left_commute];
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   191
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   192
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   193
(*** Difference ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   194
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   195
Goal "m:nat ==> m #- m = 0";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   196
by (induct_tac "m" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   197
by Auto_tac;
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   198
qed "diff_self_eq_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   199
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   200
(*Addition is the inverse of subtraction*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   201
Goal "[| n le m;  m:nat |] ==> n #+ (m#-n) = m";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   202
by (ftac lt_nat_in_nat 1);
127
eec6bb9c58ea Misc modifs such as expandshort
lcp
parents: 25
diff changeset
   203
by (etac nat_succI 1);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   204
by (etac rev_mp 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   206
by (ALLGOALS Asm_simp_tac);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   207
qed "add_diff_inverse";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   208
5504
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   209
Goal "[| n le m;  m:nat |] ==> (m#-n) #+ n = m";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   210
by (ftac lt_nat_in_nat 1);
5504
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   211
by (etac nat_succI 1);
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   212
by (asm_simp_tac (simpset() addsimps [add_commute, add_diff_inverse]) 1);
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   213
qed "add_diff_inverse2";
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   214
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   215
(*Proof is IDENTICAL to that above*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   216
Goal "[| n le m;  m:nat |] ==> succ(m) #- n = succ(m#-n)";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   217
by (ftac lt_nat_in_nat 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   218
by (etac nat_succI 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   219
by (etac rev_mp 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   220
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   221
by (ALLGOALS Asm_simp_tac);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   222
qed "diff_succ";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   223
5341
eb105c6931a4 new theorem zero_less_diff
paulson
parents: 5325
diff changeset
   224
Goal "[| m: nat; n: nat |] ==> 0 < n #- m  <->  m<n";
eb105c6931a4 new theorem zero_less_diff
paulson
parents: 5325
diff changeset
   225
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
eb105c6931a4 new theorem zero_less_diff
paulson
parents: 5325
diff changeset
   226
by (ALLGOALS Asm_simp_tac);
eb105c6931a4 new theorem zero_less_diff
paulson
parents: 5325
diff changeset
   227
qed "zero_less_diff";
eb105c6931a4 new theorem zero_less_diff
paulson
parents: 5325
diff changeset
   228
Addsimps [zero_less_diff];
eb105c6931a4 new theorem zero_less_diff
paulson
parents: 5325
diff changeset
   229
eb105c6931a4 new theorem zero_less_diff
paulson
parents: 5325
diff changeset
   230
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   231
(** Subtraction is the inverse of addition. **)
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   232
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   233
Goal "[| m:nat;  n:nat |] ==> (n#+m) #- n = m";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   234
by (induct_tac "n" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   235
by Auto_tac;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   236
qed "diff_add_inverse";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   237
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   238
Goal "[| m:nat;  n:nat |] ==> (m#+n) #- n = m";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   239
by (res_inst_tac [("m1","m")] (add_commute RS ssubst) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   240
by (REPEAT (ares_tac [diff_add_inverse] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   241
qed "diff_add_inverse2";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   242
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   243
Goal "[| k:nat; m: nat; n: nat |] ==> (k#+m) #- (k#+n) = m #- n";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   244
by (induct_tac "k" 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   245
by (ALLGOALS Asm_simp_tac);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   246
qed "diff_cancel";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   247
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   248
Goal "[| k:nat; m: nat; n: nat |] ==> (m#+k) #- (n#+k) = m #- n";
8551
5c22595bc599 tidied using new "inst" rule
paulson
parents: 8201
diff changeset
   249
by (asm_simp_tac
5c22595bc599 tidied using new "inst" rule
paulson
parents: 8201
diff changeset
   250
    (simpset() addsimps [inst "n" "k" add_commute, diff_cancel]) 1);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   251
qed "diff_cancel2";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   252
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   253
Goal "[| m:nat;  n:nat |] ==> n #- (n#+m) = 0";
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   254
by (induct_tac "n" 1);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   255
by Auto_tac;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   256
qed "diff_add_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   257
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   258
(** Difference distributes over multiplication **)
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   259
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   260
Goal "[| m:nat; n: nat; k:nat |] ==> (m #- n) #* k = (m #* k) #- (n #* k)";
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   261
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   262
by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_cancel])));
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   263
qed "diff_mult_distrib" ;
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   264
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   265
Goal "[| m:nat; n: nat; k:nat |] ==> k #* (m #- n) = (k #* m) #- (k #* n)";
8551
5c22595bc599 tidied using new "inst" rule
paulson
parents: 8201
diff changeset
   266
by (asm_simp_tac
5c22595bc599 tidied using new "inst" rule
paulson
parents: 8201
diff changeset
   267
    (simpset() addsimps [inst "m" "k" mult_commute, diff_mult_distrib]) 1);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   268
qed "diff_mult_distrib2" ;
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   269
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   270
(*** Remainder ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   271
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   272
Goal "[| 0<n;  n le m;  m:nat |] ==> m #- n < m";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   273
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   274
by (etac rev_mp 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   275
by (etac rev_mp 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   276
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   277
by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_le_self])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   278
qed "div_termination";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   279
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   280
val div_rls =   (*for mod and div*)
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   281
    nat_typechecks @
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   282
    [Ord_transrec_type, apply_type, div_termination RS ltD, if_type,
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   283
     nat_into_Ord, not_lt_iff_le RS iffD1];
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   284
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
   285
val div_ss = simpset() addsimps [div_termination RS ltD,
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   286
				 not_lt_iff_le RS iffD2];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   287
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   288
(*Type checking depends upon termination!*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   289
Goalw [mod_def] "[| 0<n;  m:nat;  n:nat |] ==> m mod n : nat";
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   290
by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   291
qed "mod_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
   292
AddTCs [mod_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   293
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   294
Goal "[| 0<n;  m<n |] ==> m mod n = m";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   295
by (rtac (mod_def RS def_transrec RS trans) 1);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   296
by (asm_simp_tac div_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   297
qed "mod_less";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   298
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   299
Goal "[| 0<n;  n le m;  m:nat |] ==> m mod n = (m#-n) mod n";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   300
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   301
by (rtac (mod_def RS def_transrec RS trans) 1);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   302
by (asm_simp_tac div_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   303
qed "mod_geq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   304
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   305
Addsimps [mod_type, mod_less, mod_geq];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   306
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   307
(*** Quotient ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   308
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   309
(*Type checking depends upon termination!*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4839
diff changeset
   310
Goalw [div_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   311
    "[| 0<n;  m:nat;  n:nat |] ==> m div n : nat";
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   312
by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   313
qed "div_type";
6153
bff90585cce5 new typechecking solver for the simplifier
paulson
parents: 6070
diff changeset
   314
AddTCs [div_type];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   315
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   316
Goal "[| 0<n;  m<n |] ==> m div n = 0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   317
by (rtac (div_def RS def_transrec RS trans) 1);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   318
by (asm_simp_tac div_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   319
qed "div_less";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   320
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   321
Goal "[| 0<n;  n le m;  m:nat |] ==> m div n = succ((m#-n) div n)";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   322
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   323
by (rtac (div_def RS def_transrec RS trans) 1);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   324
by (asm_simp_tac div_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   325
qed "div_geq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   326
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   327
Addsimps [div_type, div_less, div_geq];
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   328
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   329
(*A key result*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   330
Goal "[| 0<n;  m:nat;  n:nat |] ==> (m div n)#*n #+ m mod n = m";
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   331
by (etac complete_induct 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   332
by (excluded_middle_tac "x<n" 1);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   333
(*case x<n*)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   334
by (Asm_simp_tac 2);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   335
(*case n le x*)
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   336
by (asm_full_simp_tac
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
   337
     (simpset() addsimps [not_lt_iff_le, add_assoc,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   338
                         div_termination RS ltD, add_diff_inverse]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   339
qed "mod_div_equality";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   340
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5529
diff changeset
   341
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5529
diff changeset
   342
(*** Further facts about mod (mainly for mutilated chess board) ***)
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   343
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5529
diff changeset
   344
Goal "[| 0<n;  m:nat;  n:nat |] \
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5529
diff changeset
   345
\     ==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   346
by (etac complete_induct 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   347
by (excluded_middle_tac "succ(x)<n" 1);
1623
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   348
(* case succ(x) < n *)
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   349
by (asm_simp_tac (simpset() addsimps [nat_le_refl RS lt_trans, 
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   350
				      succ_neq_self]) 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   351
by (asm_simp_tac (simpset() addsimps [ltD RS mem_imp_not_eq]) 2);
1623
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   352
(* case n le succ(x) *)
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
   353
by (asm_full_simp_tac (simpset() addsimps [not_lt_iff_le]) 1);
1623
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   354
by (etac leE 1);
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   355
(*equality case*)
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   356
by (asm_full_simp_tac (simpset() addsimps [diff_self_eq_0]) 2);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   357
by (asm_simp_tac (simpset() addsimps [div_termination RS ltD, diff_succ]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   358
qed "mod_succ";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   359
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   360
Goal "[| 0<n;  m:nat;  n:nat |] ==> m mod n < n";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   361
by (etac complete_induct 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   362
by (excluded_middle_tac "x<n" 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   363
(*case x<n*)
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   364
by (asm_simp_tac (simpset() addsimps [mod_less]) 2);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   365
(*case n le x*)
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   366
by (asm_full_simp_tac
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
   367
     (simpset() addsimps [not_lt_iff_le, mod_geq, div_termination RS ltD]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   368
qed "mod_less_divisor";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   369
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   370
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5529
diff changeset
   371
Goal "[| k: nat; b<2 |] ==> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   372
by (subgoal_tac "k mod 2: 2" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   373
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2);
1623
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   374
by (dtac ltD 1);
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   375
by Auto_tac;
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   376
qed "mod2_cases";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   377
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   378
Goal "m:nat ==> succ(succ(m)) mod 2 = m mod 2";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   379
by (subgoal_tac "m mod 2: 2" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   380
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   381
by (asm_simp_tac (simpset() addsimps [mod_succ] setloop Step_tac) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   382
qed "mod2_succ_succ";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   383
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   384
Goal "m:nat ==> (m#+m) mod 2 = 0";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   385
by (induct_tac "m" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   386
by (simp_tac (simpset() addsimps [mod_less]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   387
by (asm_simp_tac (simpset() addsimps [mod2_succ_succ, add_succ_right]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   388
qed "mod2_add_self";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   389
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   390
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   391
(**** Additional theorems about "le" ****)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   392
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   393
Goal "[| m:nat;  n:nat |] ==> m le m #+ n";
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   394
by (induct_tac "m" 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   395
by (ALLGOALS Asm_simp_tac);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   396
qed "add_le_self";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   397
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   398
Goal "[| m:nat;  n:nat |] ==> m le n #+ m";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1793
diff changeset
   399
by (stac add_commute 1);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   400
by (REPEAT (ares_tac [add_le_self] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   401
qed "add_le_self2";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   402
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   403
(*** Monotonicity of Addition ***)
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   404
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   405
(*strict, in 1st argument; proof is by rule induction on 'less than'*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   406
Goal "[| i<j; j:nat; k:nat |] ==> i#+k < j#+k";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   407
by (ftac lt_nat_in_nat 1);
127
eec6bb9c58ea Misc modifs such as expandshort
lcp
parents: 25
diff changeset
   408
by (assume_tac 1);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   409
by (etac succ_lt_induct 1);
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
   410
by (ALLGOALS (asm_simp_tac (simpset() addsimps [leI])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   411
qed "add_lt_mono1";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   412
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   413
(*strict, in both arguments*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   414
Goal "[| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l";
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   415
by (rtac (add_lt_mono1 RS lt_trans) 1);
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   416
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1));
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1793
diff changeset
   417
by (EVERY [stac add_commute 1,
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1793
diff changeset
   418
           stac add_commute 3,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   419
           rtac add_lt_mono1 5]);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   420
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   421
qed "add_lt_mono";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   422
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   423
(*A [clumsy] way of lifting < monotonicity to le monotonicity *)
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5147
diff changeset
   424
val lt_mono::ford::prems = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   425
     "[| !!i j. [| i<j; j:k |] ==> f(i) < f(j); \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   426
\        !!i. i:k ==> Ord(f(i));                \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   427
\        i le j;  j:k                           \
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   428
\     |] ==> f(i) le f(j)";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   429
by (cut_facts_tac prems 1);
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2637
diff changeset
   430
by (blast_tac (le_cs addSIs [lt_mono,ford] addSEs [leE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   431
qed "Ord_lt_mono_imp_le_mono";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   432
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   433
(*le monotonicity, 1st argument*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   434
Goal "[| i le j; j:nat; k:nat |] ==> i#+k le j#+k";
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3736
diff changeset
   435
by (res_inst_tac [("f", "%j. j#+k")] Ord_lt_mono_imp_le_mono 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 127
diff changeset
   436
by (REPEAT (ares_tac [add_lt_mono1, add_type RS nat_into_Ord] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   437
qed "add_le_mono1";
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   438
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   439
(* le monotonicity, BOTH arguments*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   440
Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l";
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   441
by (rtac (add_le_mono1 RS le_trans) 1);
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   442
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1));
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1793
diff changeset
   443
by (EVERY [stac add_commute 1,
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1793
diff changeset
   444
           stac add_commute 3,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 760
diff changeset
   445
           rtac add_le_mono1 5]);
25
3ac1c0c0016e ordinal: DEFINITION of < and le to replace : and <= on ordinals! Many
lcp
parents: 14
diff changeset
   446
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 437
diff changeset
   447
qed "add_le_mono";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   448
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   449
(*** Monotonicity of Multiplication ***)
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   450
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   451
Goal "[| i le j; j:nat; k:nat |] ==> i#*k le j#*k";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   452
by (ftac lt_nat_in_nat 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   453
by (induct_tac "k" 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   454
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_le_mono])));
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   455
qed "mult_le_mono1";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   456
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   457
(* le monotonicity, BOTH arguments*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   458
Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#*k le j#*l";
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   459
by (rtac (mult_le_mono1 RS le_trans) 1);
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   460
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1));
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1793
diff changeset
   461
by (EVERY [stac mult_commute 1,
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1793
diff changeset
   462
           stac mult_commute 3,
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   463
           rtac mult_le_mono1 5]);
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   464
by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1));
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   465
qed "mult_le_mono";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   466
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   467
(*strict, in 1st argument; proof is by induction on k>0*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   468
Goal "[| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j";
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   469
by (etac zero_lt_natE 1);
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   470
by (ftac lt_nat_in_nat 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   471
by (ALLGOALS Asm_simp_tac);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   472
by (induct_tac "x" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   473
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_lt_mono])));
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   474
qed "mult_lt_mono2";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   475
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   476
Goal "[| i<j; 0<c; i:nat; j:nat; c:nat |] ==> i#*c < j#*c";
4839
a7322db15065 new thm mult_lt_mono1
paulson
parents: 4091
diff changeset
   477
by (asm_simp_tac (simpset() addsimps [mult_lt_mono2, mult_commute]) 1);
a7322db15065 new thm mult_lt_mono1
paulson
parents: 4091
diff changeset
   478
qed "mult_lt_mono1";
a7322db15065 new thm mult_lt_mono1
paulson
parents: 4091
diff changeset
   479
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   480
Goal "[| m: nat; n: nat |] ==> 0 < m#*n <-> 0<m & 0<n";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   481
by (best_tac (claset() addEs [natE] addss (simpset())) 1);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   482
qed "zero_lt_mult_iff";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   483
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   484
Goal "[| m: nat; n: nat |] ==> m#*n = 1 <-> m=1 & n=1";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   485
by (best_tac (claset() addEs [natE] addss (simpset())) 1);
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   486
qed "mult_eq_1_iff";
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   487
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   488
(*Cancellation law for division*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   489
Goal "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n";
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   490
by (eres_inst_tac [("i","m")] complete_induct 1);
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   491
by (excluded_middle_tac "x<n" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   492
by (asm_simp_tac (simpset() addsimps [div_less, zero_lt_mult_iff, 
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   493
                                     mult_lt_mono2]) 2);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   494
by (asm_full_simp_tac
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
   495
     (simpset() addsimps [not_lt_iff_le, 
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   496
                         zero_lt_mult_iff, le_refl RS mult_le_mono, div_geq,
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   497
                         diff_mult_distrib2 RS sym,
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   498
                         div_termination RS ltD]) 1);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   499
qed "div_cancel";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   500
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   501
Goal "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> \
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   502
\        (k#*m) mod (k#*n) = k #* (m mod n)";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   503
by (eres_inst_tac [("i","m")] complete_induct 1);
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   504
by (excluded_middle_tac "x<n" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   505
by (asm_simp_tac (simpset() addsimps [mod_less, zero_lt_mult_iff, 
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   506
                                     mult_lt_mono2]) 2);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   507
by (asm_full_simp_tac
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
   508
     (simpset() addsimps [not_lt_iff_le, 
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   509
                         zero_lt_mult_iff, le_refl RS mult_le_mono, mod_geq,
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   510
                         diff_mult_distrib2 RS sym,
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   511
                         div_termination RS ltD]) 1);
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   512
qed "mult_mod_distrib";
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   513
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   514
(*Lemma for gcd*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   515
Goal "[| m = m#*n; m: nat; n: nat |] ==> n=1 | m=0";
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   516
by (rtac disjCI 1);
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   517
by (dtac sym 1);
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   518
by (rtac Ord_linear_lt 1 THEN REPEAT_SOME (ares_tac [nat_into_Ord,nat_1I]));
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   519
by (dtac (nat_into_Ord RS Ord_0_lt RSN (2,mult_lt_mono2)) 2);
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   520
by Auto_tac;
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1708
diff changeset
   521
qed "mult_eq_self_implies_10";
1708
8f782b919043 tidied some proofs
paulson
parents: 1623
diff changeset
   522
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   523
(*Thanks to Sten Agerholm*)
5504
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   524
Goal "[|m#+n le m#+k; m:nat; n:nat; k:nat|] ==> n le k";
2493
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   525
by (etac rev_mp 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   526
by (induct_tac "m" 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   527
by (Asm_simp_tac 1);
3736
39ee3d31cfbc Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents: 3207
diff changeset
   528
by Safe_tac;
8201
a81d18b0a9b1 tidied some proofs
paulson
parents: 8127
diff changeset
   529
by (asm_full_simp_tac (simpset() addsimps [not_le_iff_lt]) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   530
qed "add_le_elim1";
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
   531
5504
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   532
Goal "[| m<n; n: nat |] ==> EX k: nat. n = succ(m#+k)";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6163
diff changeset
   533
by (ftac lt_nat_in_nat 1 THEN assume_tac 1);
6163
be8234f37e48 expandshort
paulson
parents: 6153
diff changeset
   534
by (etac rev_mp 1);
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6068
diff changeset
   535
by (induct_tac "n" 1);
5504
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   536
by (ALLGOALS (simp_tac (simpset() addsimps [le_iff])));
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   537
by (blast_tac (claset() addSEs [leE] 
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   538
                        addSIs [add_0_right RS sym, add_succ_right RS sym]) 1);
739b777e4355 new theorem less_imp_Suc_add
paulson
parents: 5341
diff changeset
   539
qed_spec_mp "less_imp_Suc_add";