src/HOL/Integ/IntArith.ML
author kleing
Tue, 13 May 2003 08:59:21 +0200
changeset 14024 213dcc39358f
parent 13837 8dd150d36c65
permissions -rw-r--r--
HOL-Real -> HOL-Complex
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents: 9269
diff changeset
     1
(*  Title:      HOL/Integ/IntArith.ML
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
     3
    Authors:    Larry Paulson and Tobias Nipkow
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
     4
*)
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
     5
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
     6
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11868
diff changeset
     7
Goal "x - - y = x + (y::int)";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11868
diff changeset
     8
by (Simp_tac 1); 
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11868
diff changeset
     9
qed "int_diff_minus_eq";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11868
diff changeset
    10
Addsimps [int_diff_minus_eq];
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11868
diff changeset
    11
9269
b62d5265b959 added zabs to arith_tac
nipkow
parents: 9214
diff changeset
    12
Goal "abs(abs(x::int)) = abs(x)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    13
by (arith_tac 1);
9214
9454f30eacc7 Defined abs on int.
nipkow
parents: 9079
diff changeset
    14
qed "abs_abs";
9454f30eacc7 Defined abs on int.
nipkow
parents: 9079
diff changeset
    15
Addsimps [abs_abs];
9454f30eacc7 Defined abs on int.
nipkow
parents: 9079
diff changeset
    16
9269
b62d5265b959 added zabs to arith_tac
nipkow
parents: 9214
diff changeset
    17
Goal "abs(-(x::int)) = abs(x)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    18
by (arith_tac 1);
9214
9454f30eacc7 Defined abs on int.
nipkow
parents: 9079
diff changeset
    19
qed "abs_minus";
9454f30eacc7 Defined abs on int.
nipkow
parents: 9079
diff changeset
    20
Addsimps [abs_minus];
9454f30eacc7 Defined abs on int.
nipkow
parents: 9079
diff changeset
    21
9269
b62d5265b959 added zabs to arith_tac
nipkow
parents: 9214
diff changeset
    22
Goal "abs(x+y) <= abs(x) + abs(y::int)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    23
by (arith_tac 1);
9269
b62d5265b959 added zabs to arith_tac
nipkow
parents: 9214
diff changeset
    24
qed "triangle_ineq";
b62d5265b959 added zabs to arith_tac
nipkow
parents: 9214
diff changeset
    25
9214
9454f30eacc7 Defined abs on int.
nipkow
parents: 9079
diff changeset
    26
10228
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    27
(*** Intermediate value theorems ***)
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    28
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
    29
Goal "(ALL i<n::nat. abs(f(i+1) - f i) <= 1) --> \
10228
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    30
\     f 0 <= k --> k <= f n --> (EX i <= n. f i = (k::int))";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    31
by (induct_tac "n" 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    32
 by (Asm_simp_tac 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    33
by (strip_tac 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    34
by (etac impE 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    35
 by (Asm_full_simp_tac 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    36
by (eres_inst_tac [("x","n")] allE 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    37
by (Asm_full_simp_tac 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    38
by (case_tac "k = f(n+1)" 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    39
 by (Force_tac 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    40
by (etac impE 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    41
 by (asm_full_simp_tac (simpset() addsimps [zabs_def] 
10702
9e6898befad4 new simprule zero_less_abs_iff
paulson
parents: 10646
diff changeset
    42
                                 addsplits [split_if_asm]) 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    43
by (blast_tac (claset() addIs [le_SucI]) 1);
10228
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    44
val lemma = result();
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    45
11770
b6bb7a853dd2 moved rulify to ObjectLogic;
wenzelm
parents: 11704
diff changeset
    46
bind_thm("nat0_intermed_int_val", ObjectLogic.rulify_no_asm lemma);
10228
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    47
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
    48
Goal "[| !i. m <= i & i < n --> abs(f(i + 1::nat) - f i) <= 1; m < n; \
10228
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    49
\        f m <= k; k <= f n |] ==> ? i. m <= i & i <= n & f i = (k::int)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    50
by (cut_inst_tac [("n","n-m"),("f", "%i. f(i+m)"),("k","k")]lemma 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    51
by (Asm_full_simp_tac 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    52
by (etac impE 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    53
 by (strip_tac 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    54
 by (eres_inst_tac [("x","i+m")] allE 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    55
 by (arith_tac 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    56
by (etac exE 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    57
by (res_inst_tac [("x","i+m")] exI 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12018
diff changeset
    58
by (arith_tac 1);
10228
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    59
qed "nat_intermed_int_val";
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    60
e653cb933293 added intermediate value thms
nipkow
parents: 9633
diff changeset
    61
9063
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    62
(*** Some convenient biconditionals for products of signs ***)
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
    63
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
    64
Goal "[| (0::int) < i; 0 < j |] ==> 0 < i*j";
9063
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    65
by (dtac zmult_zless_mono1 1);
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    66
by Auto_tac; 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    67
qed "zmult_pos";
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
    68
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
    69
Goal "[| i < (0::int); j < 0 |] ==> 0 < i*j";
9063
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    70
by (dtac zmult_zless_mono1_neg 1);
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    71
by Auto_tac; 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    72
qed "zmult_neg";
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
    73
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
    74
Goal "[| (0::int) < i; j < 0 |] ==> i*j < 0";
9063
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    75
by (dtac zmult_zless_mono1_neg 1);
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    76
by Auto_tac; 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    77
qed "zmult_pos_neg";
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
    78
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
    79
Goal "((0::int) < x*y) = (0 < x & 0 < y | x < 0 & y < 0)";
9063
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    80
by (auto_tac (claset(), 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    81
              simpset() addsimps [order_le_less, linorder_not_less,
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    82
	                          zmult_pos, zmult_neg]));
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    83
by (ALLGOALS (rtac ccontr)); 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    84
by (auto_tac (claset(), 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    85
	      simpset() addsimps [order_le_less, linorder_not_less]));
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    86
by (ALLGOALS (etac rev_mp)); 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    87
by (ALLGOALS (dtac zmult_pos_neg THEN' assume_tac));
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    88
by (auto_tac (claset() addDs [order_less_not_sym], 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    89
              simpset() addsimps [zmult_commute]));  
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    90
qed "int_0_less_mult_iff";
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
    91
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
    92
Goal "((0::int) <= x*y) = (0 <= x & 0 <= y | x <= 0 & y <= 0)";
9063
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    93
by (auto_tac (claset(), 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    94
              simpset() addsimps [order_le_less, linorder_not_less,  
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    95
                                  int_0_less_mult_iff]));
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    96
qed "int_0_le_mult_iff";
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
    97
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
    98
Goal "(x*y < (0::int)) = (0 < x & y < 0 | x < 0 & 0 < y)";
9063
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
    99
by (auto_tac (claset(), 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   100
              simpset() addsimps [int_0_le_mult_iff, 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   101
                                  linorder_not_le RS sym]));
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   102
by (auto_tac (claset() addDs [order_less_not_sym],  
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   103
              simpset() addsimps [linorder_not_le]));
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   104
qed "zmult_less_0_iff";
7707
1f4b67fdfdae simprocs now in IntArith;
wenzelm
parents:
diff changeset
   105
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   106
Goal "(x*y <= (0::int)) = (0 <= x & y <= 0 | x <= 0 & 0 <= y)";
9063
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   107
by (auto_tac (claset() addDs [order_less_not_sym], 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   108
              simpset() addsimps [int_0_less_mult_iff, 
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   109
                                  linorder_not_less RS sym]));
0d7628966069 new lemmas for signs of products
paulson
parents: 9000
diff changeset
   110
qed "zmult_le_0_iff";
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   111
10476
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   112
Goal "abs (x * y) = abs x * abs (y::int)";
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   113
by (simp_tac (simpset () delsimps [thm "number_of_reorient"] addsplits [zabs_split] 
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   114
                         addsplits [zabs_split] 
10702
9e6898befad4 new simprule zero_less_abs_iff
paulson
parents: 10646
diff changeset
   115
                         addsimps [zmult_less_0_iff, zle_def]) 1);
10476
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   116
qed "abs_mult";
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   117
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   118
Goal "(abs x = 0) = (x = (0::int))";
10476
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   119
by (simp_tac (simpset () addsplits [zabs_split]) 1);
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   120
qed "abs_eq_0";
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   121
AddIffs [abs_eq_0];
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   122
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   123
Goal "(0 < abs x) = (x ~= (0::int))";
10702
9e6898befad4 new simprule zero_less_abs_iff
paulson
parents: 10646
diff changeset
   124
by (simp_tac (simpset () addsplits [zabs_split]) 1);
9e6898befad4 new simprule zero_less_abs_iff
paulson
parents: 10646
diff changeset
   125
by (arith_tac 1);
9e6898befad4 new simprule zero_less_abs_iff
paulson
parents: 10646
diff changeset
   126
qed "zero_less_abs_iff";
9e6898befad4 new simprule zero_less_abs_iff
paulson
parents: 10646
diff changeset
   127
AddIffs [zero_less_abs_iff];
9e6898befad4 new simprule zero_less_abs_iff
paulson
parents: 10646
diff changeset
   128
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   129
Goal "0 <= x * (x::int)";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   130
by (subgoal_tac "(- x) * x <= 0" 1);
10476
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   131
 by (Asm_full_simp_tac 1);
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   132
by (simp_tac (HOL_basic_ss addsimps [zmult_le_0_iff]) 1);
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   133
by Auto_tac;
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   134
qed "square_nonzero";
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   135
Addsimps [square_nonzero];
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   136
AddIs [square_nonzero];
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   137
dbc181a32702 added abs_mult, abs_eq_0, square_nonzero;
wenzelm
parents: 10228
diff changeset
   138
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   139
(*** Products and 1, by T. M. Rasmussen ***)
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   140
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   141
Goal "(m = m*(n::int)) = (n = 1 | m = 0)";
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   142
by Auto_tac;
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   143
by (subgoal_tac "m*1 = m*n" 1);
9633
a71a83253997 better rules for cancellation of common factors across comparisons
paulson
parents: 9509
diff changeset
   144
by (dtac (zmult_cancel1 RS iffD1) 1); 
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   145
by Auto_tac;
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   146
qed "zmult_eq_self_iff";
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   147
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   148
Goal "[| 1 < m; 1 < n |] ==> 1 < m*(n::int)";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   149
by (res_inst_tac [("y","1*n")] order_less_trans 1);
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   150
by (rtac zmult_zless_mono1 2);
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   151
by (ALLGOALS Asm_simp_tac);
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   152
qed "zless_1_zmult";
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   153
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   154
Goal "[| 0 < n; n ~= 1 |] ==> 1 < (n::int)";
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   155
by (arith_tac 1);
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   156
val lemma = result();
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   157
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   158
Goal "0 < (m::int) ==> (m * n = 1) = (m = 1 & n = 1)";
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   159
by Auto_tac;
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   160
by (case_tac "m=1" 1);
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   161
by (case_tac "n=1" 2);
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   162
by (case_tac "m=1" 4);
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   163
by (case_tac "n=1" 5);
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   164
by Auto_tac;
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   165
by distinct_subgoals_tac;
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   166
by (subgoal_tac "1<m*n" 1);
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   167
by (Asm_full_simp_tac 1);
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   168
by (rtac zless_1_zmult 1);
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   169
by (ALLGOALS (rtac lemma));
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   170
by Auto_tac;  
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   171
by (subgoal_tac "0<m*n" 1);
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   172
by (Asm_simp_tac 2);
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   173
by (dtac (int_0_less_mult_iff RS iffD1) 1);
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   174
by Auto_tac;  
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   175
qed "pos_zmult_eq_1_iff";
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   176
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   177
Goal "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   178
by (case_tac "0<m" 1);
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   179
by (asm_simp_tac (simpset() addsimps [pos_zmult_eq_1_iff]) 1);
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   180
by (case_tac "m=0" 1);
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   181
by (asm_simp_tac (simpset () delsimps [thm "number_of_reorient"]) 1);
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11770
diff changeset
   182
by (subgoal_tac "0 < -m" 1);
9509
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   183
by (arith_tac 2);
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   184
by (dres_inst_tac [("n","-n")] pos_zmult_eq_1_iff 1); 
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   185
by Auto_tac;  
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   186
qed "zmult_eq_1_iff";
0f3ee1f89ca8 introduction of integer exponentiation
paulson
parents: 9436
diff changeset
   187
13837
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   188
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   189
(*** More about nat ***)
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   190
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   191
Goal "[| (0::int) <= z;  0 <= z' |] ==> nat (z+z') = nat z + nat z'";
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   192
by (rtac (inj_int RS injD) 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   193
by (asm_simp_tac (simpset() addsimps [zadd_int RS sym]) 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   194
qed "nat_add_distrib";
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   195
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   196
Goal "[| (0::int) <= z';  z' <= z |] ==> nat (z-z') = nat z - nat z'";
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   197
by (rtac (inj_int RS injD) 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   198
by (asm_simp_tac (simpset() addsimps [zdiff_int RS sym, nat_le_eq_zle]) 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   199
qed "nat_diff_distrib";
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   200
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   201
Goal "(0::int) <= z ==> nat (z*z') = nat z * nat z'";
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   202
by (case_tac "0 <= z'" 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   203
by (asm_full_simp_tac (simpset() addsimps [zmult_le_0_iff]) 2);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   204
by (rtac (inj_int RS injD) 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   205
by (asm_simp_tac (simpset() addsimps [zmult_int RS sym,
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   206
				      int_0_le_mult_iff]) 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   207
qed "nat_mult_distrib";
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   208
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   209
Goal "z <= (0::int) ==> nat(z*z') = nat(-z) * nat(-z')"; 
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   210
by (rtac trans 1); 
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   211
by (rtac nat_mult_distrib 2); 
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   212
by Auto_tac;  
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   213
qed "nat_mult_distrib_neg";
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   214
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   215
Goal "nat (abs (w * z)) = nat (abs w) * nat (abs z)";
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   216
by (case_tac "z=0 | w=0" 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   217
by Auto_tac;  
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   218
by (simp_tac (simpset() addsimps [zabs_def, nat_mult_distrib RS sym, 
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   219
                          nat_mult_distrib_neg RS sym, zmult_less_0_iff]) 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   220
by (arith_tac 1);
8dd150d36c65 Reorganized, moving many results about the integer dvd relation from IntPrimes
paulson
parents: 13187
diff changeset
   221
qed "nat_abs_mult_distrib";