src/ZF/ArithSimp.ML
author wenzelm
Thu, 15 Nov 2001 18:20:13 +0100
changeset 12207 4dff931b852f
parent 12089 34e7693271a9
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:
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     1
(*  Title:      ZF/ArithSimp.ML
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     2
    ID:         $Id$
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     4
    Copyright   2000  University of Cambridge
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     5
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     6
Arithmetic with simplification
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     7
*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     8
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
     9
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    10
Addsimprocs ArithData.nat_cancel;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    11
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    12
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    13
(*** Difference ***)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    14
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    15
Goal "m #- m = 0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    16
by (subgoal_tac "natify(m) #- natify(m) = 0" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    17
by (rtac (natify_in_nat RS nat_induct) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    18
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    19
qed "diff_self_eq_0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    20
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    21
(**Addition is the inverse of subtraction**)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    22
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    23
(*We need m:nat even if we replace the RHS by natify(m), for consider e.g.
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    24
  n=2, m=omega; then n + (m-n) = 2 + (0-2) = 2 ~= 0 = natify(m).*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    25
Goal "[| n le m;  m:nat |] ==> n #+ (m#-n) = m";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    26
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    27
by (etac rev_mp 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    28
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    29
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    30
qed "add_diff_inverse";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    31
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    32
Goal "[| n le m;  m:nat |] ==> (m#-n) #+ n = m";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    33
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    34
by (asm_simp_tac (simpset() addsimps [add_commute, add_diff_inverse]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    35
qed "add_diff_inverse2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    36
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    37
(*Proof is IDENTICAL to that of add_diff_inverse*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    38
Goal "[| n le m;  m:nat |] ==> succ(m) #- n = succ(m#-n)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    39
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    40
by (etac rev_mp 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    41
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    42
by (ALLGOALS Asm_simp_tac);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    43
qed "diff_succ";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    44
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    45
Goal "[| m: nat; n: nat |] ==> 0 < (n #- m)   <->   m<n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    46
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    47
by (ALLGOALS Asm_simp_tac);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    48
qed "zero_less_diff";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    49
Addsimps [zero_less_diff];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    50
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    51
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    52
(** Difference distributes over multiplication **)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    53
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    54
Goal "(m #- n) #* k = (m #* k) #- (n #* k)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    55
by (subgoal_tac "(natify(m) #- natify(n)) #* natify(k) = \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    56
\                (natify(m) #* natify(k)) #- (natify(n) #* natify(k))" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    57
by (res_inst_tac [("m","natify(m)"),("n","natify(n)")] diff_induct 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    58
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [diff_cancel])));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    59
qed "diff_mult_distrib" ;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    60
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    61
Goal "k #* (m #- n) = (k #* m) #- (k #* n)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    62
by (simp_tac
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    63
    (simpset() addsimps [inst "m" "k" mult_commute, diff_mult_distrib]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    64
qed "diff_mult_distrib2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    65
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    66
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    67
(*** Remainder ***)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    68
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    69
(*We need m:nat even with natify*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    70
Goal "[| 0<n;  n le m;  m:nat |] ==> m #- n < m";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    71
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    72
by (etac rev_mp 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    73
by (etac rev_mp 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    74
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    75
by (ALLGOALS (asm_simp_tac (simpset() addsimps [diff_le_self])));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    76
qed "div_termination";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    77
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9883
diff changeset
    78
bind_thms ("div_rls",   (*for mod and div*)
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    79
    nat_typechecks @
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    80
    [Ord_transrec_type, apply_funtype, div_termination RS ltD,
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9883
diff changeset
    81
     nat_into_Ord, not_lt_iff_le RS iffD1]);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    82
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    83
val div_ss = simpset() addsimps [div_termination RS ltD,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    84
				 not_lt_iff_le RS iffD2];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    85
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    86
Goalw [raw_mod_def] "[| m:nat;  n:nat |] ==> raw_mod (m, n) : nat";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    87
by (rtac Ord_transrec_type 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    88
by (auto_tac(claset(), simpset() addsimps [nat_into_Ord RS Ord_0_lt_iff]));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    89
by (REPEAT (ares_tac div_rls 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    90
qed "raw_mod_type";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    91
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    92
Goalw [mod_def] "m mod n : nat";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    93
by (simp_tac (simpset() addsimps [mod_def, raw_mod_type]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    94
qed "mod_type";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    95
AddTCs [mod_type];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    96
AddIffs [mod_type];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    97
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    98
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
    99
(** Aribtrary definitions for division by zero.  Useful to simplify 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   100
    certain equations **)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   101
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   102
Goalw [div_def] "a div 0 = 0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   103
by (rtac (raw_div_def RS def_transrec RS trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   104
by (Asm_simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   105
qed "DIVISION_BY_ZERO_DIV";  (*NOT for adding to default simpset*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   106
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   107
Goalw [mod_def] "a mod 0 = natify(a)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   108
by (rtac (raw_mod_def RS def_transrec RS trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   109
by (Asm_simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   110
qed "DIVISION_BY_ZERO_MOD";  (*NOT for adding to default simpset*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   111
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   112
fun div_undefined_case_tac s i =
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   113
  case_tac s i THEN 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   114
  asm_full_simp_tac
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   115
         (simpset() addsimps [nat_into_Ord RS Ord_0_lt_iff]) (i+1) THEN
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   116
  asm_full_simp_tac (simpset() addsimps [DIVISION_BY_ZERO_DIV, 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   117
		 	 	         DIVISION_BY_ZERO_MOD]) i;
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   118
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   119
Goal "m<n ==> raw_mod (m,n) = m";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   120
by (rtac (raw_mod_def RS def_transrec RS trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   121
by (asm_simp_tac (simpset() addsimps [div_termination RS ltD]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   122
qed "raw_mod_less";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   123
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   124
Goal "[| m<n; n : nat |] ==> m mod n = m";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   125
by (ftac lt_nat_in_nat 1 THEN assume_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   126
by (asm_simp_tac (simpset() addsimps [mod_def, raw_mod_less]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   127
qed "mod_less";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   128
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   129
Goal "[| 0<n; n le m;  m:nat |] ==> raw_mod (m, n) = raw_mod (m#-n, n)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   130
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   131
by (rtac (raw_mod_def RS def_transrec RS trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   132
by (asm_simp_tac div_ss 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   133
by (Blast_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   134
qed "raw_mod_geq";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   135
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   136
Goal "[| n le m;  m:nat |] ==> m mod n = (m#-n) mod n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   137
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   138
by (div_undefined_case_tac "n=0" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   139
by (asm_simp_tac (simpset() addsimps [mod_def, raw_mod_geq]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   140
qed "mod_geq";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   141
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   142
Addsimps [mod_less];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   143
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   144
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   145
(*** Division ***)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   146
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   147
Goalw [raw_div_def] "[| m:nat;  n:nat |] ==> raw_div (m, n) : nat";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   148
by (rtac Ord_transrec_type 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   149
by (auto_tac(claset(), simpset() addsimps [nat_into_Ord RS Ord_0_lt_iff]));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   150
by (REPEAT (ares_tac div_rls 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   151
qed "raw_div_type";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   152
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   153
Goalw [div_def] "m div n : nat";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   154
by (simp_tac (simpset() addsimps [div_def, raw_div_type]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   155
qed "div_type";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   156
AddTCs [div_type];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   157
AddIffs [div_type];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   158
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   159
Goal "m<n ==> raw_div (m,n) = 0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   160
by (rtac (raw_div_def RS def_transrec RS trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   161
by (asm_simp_tac (simpset() addsimps [div_termination RS ltD]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   162
qed "raw_div_less";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   163
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   164
Goal "[| m<n; n : nat |] ==> m div n = 0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   165
by (ftac lt_nat_in_nat 1 THEN assume_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   166
by (asm_simp_tac (simpset() addsimps [div_def, raw_div_less]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   167
qed "div_less";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   168
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   169
Goal "[| 0<n;  n le m;  m:nat |] ==> raw_div(m,n) = succ(raw_div(m#-n, n))";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   170
by (subgoal_tac "n ~= 0" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   171
by (Blast_tac 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   172
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   173
by (rtac (raw_div_def RS def_transrec RS trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   174
by (asm_simp_tac div_ss 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   175
qed "raw_div_geq";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   176
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   177
Goal "[| 0<n;  n le m;  m:nat |] ==> m div n = succ ((m#-n) div n)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   178
by (ftac lt_nat_in_nat 1 THEN etac nat_succI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   179
by (asm_simp_tac (simpset() addsimps [div_def, raw_div_geq]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   180
qed "div_geq";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   181
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   182
Addsimps [div_less, div_geq];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   183
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   184
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   185
(*A key result*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   186
Goal "[| m: nat;  n: nat |] ==> (m div n)#*n #+ m mod n = m";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   187
by (div_undefined_case_tac "n=0" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   188
by (etac complete_induct 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   189
by (case_tac "x<n" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   190
(*case n le x*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   191
by (asm_full_simp_tac
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   192
     (simpset() addsimps [not_lt_iff_le, add_assoc, mod_geq,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   193
                         div_termination RS ltD, add_diff_inverse]) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   194
(*case x<n*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   195
by (Asm_simp_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   196
val lemma = result();
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   197
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   198
Goal "(m div n)#*n #+ m mod n = natify(m)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   199
by (subgoal_tac
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   200
    "(natify(m) div natify(n))#*natify(n) #+ natify(m) mod natify(n) = \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   201
\    natify(m)" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   202
by (stac lemma 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   203
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   204
qed "mod_div_equality_natify";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   205
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   206
Goal "m: nat ==> (m div n)#*n #+ m mod n = m";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   207
by (asm_simp_tac (simpset() addsimps [mod_div_equality_natify]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   208
qed "mod_div_equality";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   209
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   210
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   211
(*** Further facts about mod (mainly for mutilated chess board) ***)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   212
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   213
Goal "[| 0<n;  m:nat;  n:nat |] \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   214
\     ==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   215
by (etac complete_induct 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   216
by (excluded_middle_tac "succ(x)<n" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   217
(* case succ(x) < n *)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   218
by (asm_simp_tac (simpset() addsimps [nat_le_refl RS lt_trans, 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   219
				      succ_neq_self]) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   220
by (asm_simp_tac (simpset() addsimps [ltD RS mem_imp_not_eq]) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   221
(* case n le succ(x) *)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   222
by (asm_full_simp_tac (simpset() addsimps [mod_geq, not_lt_iff_le]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   223
by (etac leE 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   224
(*equality case*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   225
by (asm_full_simp_tac (simpset() addsimps [diff_self_eq_0]) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   226
by (asm_simp_tac (simpset() addsimps [mod_geq, div_termination RS ltD, 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   227
				      diff_succ]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   228
val lemma = result();
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   229
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   230
Goal "n:nat ==> \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   231
\     succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   232
by (case_tac "n=0" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   233
by (asm_simp_tac (simpset() addsimps [natify_succ, DIVISION_BY_ZERO_MOD]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   234
by (subgoal_tac
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   235
    "natify(succ(m)) mod n = \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   236
\      (if succ(natify(m) mod n) = n then 0 else succ(natify(m) mod n))" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   237
by (stac natify_succ 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   238
by (rtac lemma 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   239
by (auto_tac(claset(), 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   240
	     simpset() delsimps [natify_succ] 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   241
             addsimps [nat_into_Ord RS Ord_0_lt_iff]));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   242
qed "mod_succ";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   243
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   244
Goal "[| 0<n;  n:nat |] ==> m mod n < n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   245
by (subgoal_tac "natify(m) mod n < n" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   246
by (res_inst_tac [("i","natify(m)")] complete_induct 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   247
by (case_tac "x<n" 3);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   248
(*case x<n*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   249
by (Asm_simp_tac 3);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   250
(*case n le x*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   251
by (asm_full_simp_tac
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   252
     (simpset() addsimps [mod_geq, not_lt_iff_le, div_termination RS ltD]) 3);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   253
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   254
qed "mod_less_divisor";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   255
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   256
Goal "m mod 1 = 0";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   257
by (cut_inst_tac [("n","1")] mod_less_divisor 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   258
by Auto_tac; 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   259
qed "mod_1_eq";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   260
Addsimps [mod_1_eq]; 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   261
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   262
Goal "b<2 ==> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   263
by (subgoal_tac "k mod 2: 2" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   264
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   265
by (dtac ltD 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   266
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   267
qed "mod2_cases";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   268
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   269
Goal "succ(succ(m)) mod 2 = m mod 2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   270
by (subgoal_tac "m mod 2: 2" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   271
by (asm_simp_tac (simpset() addsimps [mod_less_divisor RS ltD]) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   272
by (auto_tac (claset(), simpset() addsimps [mod_succ]));  
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   273
qed "mod2_succ_succ";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   274
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   275
Addsimps [mod2_succ_succ];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   276
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   277
Goal "(m#+m#+n) mod 2 = n mod 2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   278
by (subgoal_tac "(natify(m)#+natify(m)#+n) mod 2 = n mod 2" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   279
by (res_inst_tac [("n","natify(m)")] nat_induct 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   280
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   281
qed "mod2_add_more";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   282
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   283
Goal "(m#+m) mod 2 = 0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   284
by (cut_inst_tac [("n","0")] mod2_add_more 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   285
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   286
qed "mod2_add_self";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   287
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   288
Addsimps [mod2_add_more, mod2_add_self];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   289
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   290
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   291
(**** Additional theorems about "le" ****)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   292
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   293
Goal "m:nat ==> m le (m #+ n)";
9873
paulson
parents: 9648
diff changeset
   294
by (Asm_simp_tac 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   295
qed "add_le_self";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   296
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   297
Goal "m:nat ==> m le (n #+ m)";
9873
paulson
parents: 9648
diff changeset
   298
by (Asm_simp_tac 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   299
qed "add_le_self2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   300
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   301
(*** Monotonicity of Multiplication ***)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   302
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   303
Goal "[| i le j; j:nat |] ==> (i#*k) le (j#*k)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   304
by (subgoal_tac "natify(i)#*natify(k) le j#*natify(k)" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   305
by (ftac lt_nat_in_nat 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   306
by (res_inst_tac [("n","natify(k)")] nat_induct 3);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   307
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_le_mono])));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   308
qed "mult_le_mono1";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   309
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   310
(* le monotonicity, BOTH arguments*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   311
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:
diff changeset
   312
by (rtac (mult_le_mono1 RS le_trans) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   313
by (REPEAT (assume_tac 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   314
by (EVERY [stac mult_commute 1,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   315
           stac mult_commute 1,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   316
           rtac mult_le_mono1 1]);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   317
by (REPEAT (assume_tac 1));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   318
qed "mult_le_mono";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   319
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   320
(*strict, in 1st argument; proof is by induction on k>0.
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   321
  I can't see how to relax the typing conditions.*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   322
Goal "[| i<j; 0<k; j:nat; k:nat |] ==> k#*i < k#*j";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   323
by (etac zero_lt_natE 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   324
by (ftac lt_nat_in_nat 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   325
by (ALLGOALS Asm_simp_tac);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   326
by (induct_tac "x" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   327
by (ALLGOALS (asm_simp_tac (simpset() addsimps [add_lt_mono])));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   328
qed "mult_lt_mono2";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   329
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   330
Goal "[| i<j; 0<k; j:nat; k:nat |] ==> i#*k < j#*k";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   331
by (asm_simp_tac (simpset() addsimps [mult_lt_mono2, 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   332
				      inst "n" "k" mult_commute]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   333
qed "mult_lt_mono1";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   334
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   335
Goal "m#+n = 0 <-> natify(m)=0 & natify(n)=0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   336
by (subgoal_tac "natify(m) #+ natify(n) = 0 <-> natify(m)=0 & natify(n)=0" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   337
by (res_inst_tac [("n","natify(m)")] natE 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   338
 by (res_inst_tac [("n","natify(n)")] natE 4);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   339
by Auto_tac;  
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   340
qed "add_eq_0_iff";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   341
AddIffs [add_eq_0_iff];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   342
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   343
Goal "0 < m#*n <-> 0 < natify(m) & 0 < natify(n)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   344
by (subgoal_tac "0 < natify(m)#*natify(n) <-> \
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   345
\                0 < natify(m) & 0 < natify(n)" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   346
by (res_inst_tac [("n","natify(m)")] natE 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   347
 by (res_inst_tac [("n","natify(n)")] natE 4);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   348
  by (res_inst_tac [("n","natify(n)")] natE 3);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   349
by Auto_tac;  
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   350
qed "zero_lt_mult_iff";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   351
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   352
Goal "m#*n = 1 <-> natify(m)=1 & natify(n)=1";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   353
by (subgoal_tac "natify(m) #* natify(n) = 1 <-> natify(m)=1 & natify(n)=1" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   354
by (res_inst_tac [("n","natify(m)")] natE 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   355
 by (res_inst_tac [("n","natify(n)")] natE 4);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   356
by Auto_tac;  
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   357
qed "mult_eq_1_iff";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   358
AddIffs [zero_lt_mult_iff, mult_eq_1_iff];
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   359
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   360
Goal "[|m: nat; n: nat|] ==> (m #* n = 0) <-> (m = 0 | n = 0)";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   361
by Auto_tac;
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   362
by (etac natE 1);  
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   363
by (etac natE 2);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   364
by Auto_tac; 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   365
qed "mult_is_zero";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   366
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   367
Goal "(m #* n = 0) <-> (natify(m) = 0 | natify(n) = 0)";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   368
by (cut_inst_tac [("m","natify(m)"),("n","natify(n)")] mult_is_zero 1); 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   369
by Auto_tac; 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   370
qed "mult_is_zero_natify";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   371
AddIffs [mult_is_zero_natify];
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   372
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   373
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   374
(** Cancellation laws for common factors in comparisons **)
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   375
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   376
Goal "[| k: nat; m: nat; n: nat |] ==> (m#*k < n#*k) <-> (0<k & m<n)";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   377
by (safe_tac (claset() addSIs [mult_lt_mono1]));
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   378
by (etac natE 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   379
by Auto_tac;  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   380
by (rtac (not_le_iff_lt RS iffD1) 1); 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   381
by (dtac (not_le_iff_lt RSN (2,rev_iffD2)) 3); 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   382
by (blast_tac (claset() addIs [mult_le_mono1]) 5); 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   383
by Auto_tac;  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   384
val lemma = result();
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   385
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   386
Goal "(m#*k < n#*k) <-> (0 < natify(k) & natify(m) < natify(n))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   387
by (rtac iff_trans 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   388
by (rtac lemma 2);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   389
by Auto_tac;  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   390
qed "mult_less_cancel2";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   391
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   392
Goal "(k#*m < k#*n) <-> (0 < natify(k) & natify(m) < natify(n))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   393
by (simp_tac (simpset() addsimps [mult_less_cancel2, 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   394
                                  inst "m" "k" mult_commute]) 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   395
qed "mult_less_cancel1";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   396
Addsimps [mult_less_cancel1, mult_less_cancel2];
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   397
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   398
Goal "(m#*k le n#*k) <-> (0 < natify(k) --> natify(m) le natify(n))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   399
by (asm_simp_tac (simpset() addsimps [not_lt_iff_le RS iff_sym]) 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   400
by Auto_tac;  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   401
qed "mult_le_cancel2";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   402
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   403
Goal "(k#*m le k#*n) <-> (0 < natify(k) --> natify(m) le natify(n))";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   404
by (asm_simp_tac (simpset() addsimps [not_lt_iff_le RS iff_sym]) 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   405
by Auto_tac;  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   406
qed "mult_le_cancel1";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   407
Addsimps [mult_le_cancel1, mult_le_cancel2];
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   408
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   409
Goal "k : nat ==> k #* m le k \\<longleftrightarrow> (0 < k \\<longrightarrow> natify(m) le 1)";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   410
by (cut_inst_tac [("k","k"),("m","m"),("n","1")] mult_le_cancel1 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   411
by Auto_tac;
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   412
qed "mult_le_cancel_le1";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   413
9648
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   414
Goal "[| Ord(m); Ord(n) |] ==> m=n <-> (m le n & n le m)";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   415
by (blast_tac (claset() addIs [le_anti_sym]) 1); 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   416
qed "Ord_eq_iff_le";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   417
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   418
Goal "[| k: nat; m: nat; n: nat |] ==> (m#*k = n#*k) <-> (m=n | k=0)";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   419
by (asm_simp_tac (simpset() addsimps [inst "m" "m#*k" Ord_eq_iff_le,
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   420
                                      inst "m" "m" Ord_eq_iff_le]) 1); 
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   421
by (auto_tac (claset(), simpset() addsimps [Ord_0_lt_iff]));  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   422
val lemma = result();
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   423
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   424
Goal "(m#*k = n#*k) <-> (natify(m) = natify(n) | natify(k) = 0)";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   425
by (rtac iff_trans 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   426
by (rtac lemma 2);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   427
by Auto_tac;  
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   428
qed "mult_cancel2";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   429
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   430
Goal "(k#*m = k#*n) <-> (natify(m) = natify(n) | natify(k) = 0)";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   431
by (simp_tac (simpset() addsimps [mult_cancel2, inst "m" "k" mult_commute]) 1);
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   432
qed "mult_cancel1";
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   433
Addsimps [mult_cancel1, mult_cancel2];
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   434
35d761c7d934 better rules for cancellation of common factors across comparisons
paulson
parents: 9548
diff changeset
   435
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   436
(** Cancellation law for division **)
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   437
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   438
Goal "[| 0<n; 0<k; k:nat; m:nat; n:nat |] ==> (k#*m) div (k#*n) = m div n";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   439
by (eres_inst_tac [("i","m")] complete_induct 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   440
by (excluded_middle_tac "x<n" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   441
by (asm_simp_tac (simpset() addsimps [div_less, zero_lt_mult_iff, 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   442
                                     mult_lt_mono2]) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   443
by (asm_full_simp_tac
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   444
     (simpset() addsimps [not_lt_iff_le, 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   445
                         zero_lt_mult_iff, le_refl RS mult_le_mono, div_geq,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   446
                         diff_mult_distrib2 RS sym,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   447
                         div_termination RS ltD]) 1);
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   448
qed "div_cancel_raw";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   449
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   450
Goal "[| 0 < natify(n);  0 < natify(k) |] ==> (k#*m) div (k#*n) = m div n";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   451
by (cut_inst_tac [("k","natify(k)"),("m","natify(m)"),("n","natify(n)")] 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   452
    div_cancel_raw 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   453
by Auto_tac; 
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   454
qed "div_cancel";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   455
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   456
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   457
(** Distributive law for remainder (mod) **)
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   458
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   459
Goal "[| k:nat; m:nat; n:nat |] ==> (k#*m) mod (k#*n) = k #* (m mod n)";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   460
by (div_undefined_case_tac "k=0" 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   461
by (div_undefined_case_tac "n=0" 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   462
by (eres_inst_tac [("i","m")] complete_induct 1);
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   463
by (case_tac "x<n" 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   464
 by (asm_simp_tac
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   465
     (simpset() addsimps [mod_less, zero_lt_mult_iff, mult_lt_mono2]) 1);
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   466
by (asm_full_simp_tac
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   467
     (simpset() addsimps [not_lt_iff_le, 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   468
                         zero_lt_mult_iff, le_refl RS mult_le_mono, mod_geq,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   469
                         diff_mult_distrib2 RS sym,
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   470
                         div_termination RS ltD]) 1);
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   471
qed "mult_mod_distrib_raw";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   472
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   473
Goal "k #* (m mod n) = (k#*m) mod (k#*n)";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   474
by (cut_inst_tac [("k","natify(k)"),("m","natify(m)"),("n","natify(n)")] 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   475
    mult_mod_distrib_raw 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   476
by Auto_tac; 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   477
qed "mod_mult_distrib2";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   478
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   479
Goal "(m mod n) #* k = (m#*k) mod (n#*k)";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   480
by (simp_tac (simpset() addsimps [mult_commute, mod_mult_distrib2]) 1); 
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   481
qed "mult_mod_distrib";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   482
11386
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   483
Goal "n \\<in> nat ==> (m #+ n) mod n = m mod n";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   484
by (subgoal_tac "(n #+ m) mod n = (n #+ m #- n) mod n" 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   485
by (stac (mod_geq RS sym) 2);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   486
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_commute])));
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   487
qed "mod_add_self2_raw";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   488
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   489
Goal "(m #+ n) mod n = m mod n";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   490
by (cut_inst_tac [("n","natify(n)")] mod_add_self2_raw 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   491
by Auto_tac; 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   492
qed "mod_add_self2";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   493
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   494
Goal "(n#+m) mod n = m mod n";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   495
by (asm_simp_tac (simpset() addsimps [add_commute, mod_add_self2]) 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   496
qed "mod_add_self1";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   497
Addsimps [mod_add_self1, mod_add_self2];
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   498
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   499
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   500
Goal "k \\<in> nat ==> (m #+ k#*n) mod n = m mod n";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   501
by (etac nat_induct 1); 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   502
by (ALLGOALS
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   503
    (asm_simp_tac 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   504
     (simpset() addsimps [inst "n" "n" add_left_commute])));
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   505
qed "mod_mult_self1_raw";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   506
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   507
Goal "(m #+ k#*n) mod n = m mod n";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   508
by (cut_inst_tac [("k","natify(k)")] mod_mult_self1_raw 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   509
by Auto_tac; 
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   510
qed "mod_mult_self1";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   511
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   512
Goal "(m #+ n#*k) mod n = m mod n";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   513
by (simp_tac (simpset() addsimps [mult_commute, mod_mult_self1]) 1);
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   514
qed "mod_mult_self2";
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   515
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   516
Addsimps [mod_mult_self1, mod_mult_self2];
cf8d81cf8034 a few new and/or improved results
paulson
parents: 9907
diff changeset
   517
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   518
(*Lemma for gcd*)
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   519
Goal "m = m#*n ==> natify(n)=1 | m=0";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   520
by (subgoal_tac "m: nat" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   521
by (etac ssubst 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   522
by (rtac disjCI 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   523
by (dtac sym 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   524
by (rtac Ord_linear_lt 1 THEN REPEAT_SOME (ares_tac [nat_into_Ord,nat_1I]));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   525
by (dtac (nat_into_Ord RS Ord_0_lt RSN (2,mult_lt_mono2)) 3);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   526
by Auto_tac;
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   527
by (subgoal_tac "m #* n = 0" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   528
by (stac (mult_natify2 RS sym) 2);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   529
by (auto_tac (claset(), simpset() delsimps [mult_natify2]));  
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   530
qed "mult_eq_self_implies_10";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   531
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   532
Goal "[| m<n; n: nat |] ==> EX k: nat. n = succ(m#+k)";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   533
by (ftac lt_nat_in_nat 1 THEN assume_tac 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   534
by (etac rev_mp 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   535
by (induct_tac "n" 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   536
by (ALLGOALS (simp_tac (simpset() addsimps [le_iff])));
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   537
by (blast_tac (claset() addSEs [leE] 
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   538
                        addSIs [add_0_right RS sym, add_succ_right RS sym]) 1);
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   539
qed_spec_mp "less_imp_succ_add";
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents:
diff changeset
   540
9883
c1c8647af477 a number of new theorems
paulson
parents: 9873
diff changeset
   541
Goal "[| m: nat; n: nat |] ==> (m<n) <-> (EX k: nat. n = succ(m#+k))";
c1c8647af477 a number of new theorems
paulson
parents: 9873
diff changeset
   542
by (auto_tac (claset() addIs [less_imp_succ_add], simpset()));
c1c8647af477 a number of new theorems
paulson
parents: 9873
diff changeset
   543
qed "less_iff_succ_add";
12089
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   544
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   545
(* on nat *)
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   546
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   547
Goal "m #- n = 0 <-> natify(m) le natify(n)";
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   548
by (auto_tac (claset(), simpset() addsimps 
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   549
      [le_iff, diff_self_eq_0]));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   550
by (full_simp_tac (simpset() addsimps [less_iff_succ_add ]) 2);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   551
by (Clarify_tac 2);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   552
by (subgoal_tac "natify(m) #- natify(n) = 0" 3);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   553
by (etac subst 3);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   554
by (ALLGOALS(asm_full_simp_tac (simpset() 
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   555
             addsimps [diff_self_eq_0])));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   556
by (subgoal_tac "natify(m) #- natify(n) = 0" 2);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   557
by (etac subst 2);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   558
by (etac ssubst  3);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   559
by (rtac (not_le_iff_lt RS iffD1) 1);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   560
by (auto_tac (claset(), simpset() addsimps [le_iff]));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   561
by (subgoal_tac "natify(m) #- natify(n) = 0" 1);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   562
by (Asm_simp_tac 2);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   563
by (thin_tac "m #- n = 0" 1);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   564
by (dtac ([natify_in_nat, natify_in_nat] MRS zero_less_diff RS iffD2) 1);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   565
by Auto_tac;
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   566
qed "diff_is_0_iff";
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   567
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   568
Goal "[| a:nat; b:nat |] ==> \
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   569
\ (P(a #- b)) <-> ((a < b -->P(0)) & (ALL d:nat. a = b #+ d --> P(d)))";
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   570
by (case_tac "a le b" 1);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   571
by (subgoal_tac "natify(a) le natify(b)" 1);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   572
by (dtac (diff_is_0_iff RS iffD2) 1);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   573
by Safe_tac;
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   574
by (ALLGOALS(Asm_full_simp_tac));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   575
by (ALLGOALS(rotate_tac ~1));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   576
by (ALLGOALS(asm_full_simp_tac  (simpset() addsimps [le_iff])));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   577
by (Clarify_tac 2);
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   578
by (ALLGOALS(dtac not_lt_imp_le));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   579
by (ALLGOALS(Asm_full_simp_tac));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   580
by (ALLGOALS(dres_inst_tac [("x", "a #- b")] bspec));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   581
by (ALLGOALS(Asm_simp_tac));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   582
by (ALLGOALS(dtac add_diff_inverse));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   583
by (ALLGOALS(rotate_tac ~1));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   584
by (ALLGOALS(Asm_full_simp_tac));
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   585
qed "nat_diff_split";
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   586
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   587
34e7693271a9 Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents: 11386
diff changeset
   588