src/HOL/Integ/Int.ML
author paulson
Thu, 01 Oct 1998 18:25:56 +0200
changeset 5593 33bca87deae5
parent 5582 a356fb49e69e
child 6717 70b251dc7055
permissions -rw-r--r--
new lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Integ/Int.ML
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     2
    ID:         $Id$
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     5
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     6
Type "int" is a linear order
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     7
*)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     8
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
     9
Goal "(w<z) = neg(w-z)";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    10
by (simp_tac (simpset() addsimps [zless_def]) 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    11
qed "zless_eq_neg";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    12
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    13
Goal "(w=z) = iszero(w-z)";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    14
by (simp_tac (simpset() addsimps [iszero_def, zdiff_eq_eq]) 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    15
qed "eq_eq_iszero";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    16
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    17
Goal "(w<=z) = (~ neg(z-w))";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    18
by (simp_tac (simpset() addsimps [zle_def, zless_def]) 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    19
qed "zle_eq_not_neg";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    20
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    21
5593
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    22
(*** Inequality lemmas involving int (Suc m) ***)
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    23
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    24
Goal "(w < z + int (Suc m)) = (w < z + int m | w = z + int m)";
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    25
by (auto_tac (claset(),
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    26
	      simpset() addsimps [zless_iff_Suc_zadd, zadd_assoc, zadd_int]));
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    27
by (cut_inst_tac [("m","m")] int_Suc 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    28
by (cut_inst_tac [("m","n")] int_Suc 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    29
by (Asm_full_simp_tac 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    30
by (exhaust_tac "n" 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    31
by Auto_tac;
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    32
by (cut_inst_tac [("m","m")] int_Suc 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    33
by (full_simp_tac (simpset() addsimps zadd_ac) 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    34
by (asm_full_simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    35
by (auto_tac (claset(),
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    36
	      simpset() addsimps [zless_iff_Suc_zadd, zadd_assoc, zadd_int]));
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    37
qed "zless_add_int_Suc_eq";
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    38
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    39
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    40
Goal "(w + int (Suc m) <= z) = (w + int m < z)";
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    41
by (simp_tac (simpset() addsimps [zle_def, zless_add_int_Suc_eq]) 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    42
by (auto_tac (claset() addIs [zle_anti_sym] addEs [zless_asym],
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    43
	      simpset() addsimps [zless_imp_zle, symmetric zle_def]));
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    44
qed "add_int_Suc_zle_eq";
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    45
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    46
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    47
(* (w < int (Suc m)) = (w < int m | w = int m) *)
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    48
bind_thm ("less_int_Suc_eq",
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    49
	  simplify (simpset())
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    50
	    (read_instantiate  [("z", "int 0")] zless_add_int_Suc_eq));
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    51
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    52
Goal "(w <= int (Suc m)) = (w <= int m | w = int (Suc m))";
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    53
by (simp_tac (simpset() addsimps [less_int_Suc_eq, order_le_less]) 1);
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    54
qed "le_int_Suc_eq";
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    55
33bca87deae5 new lemmas
paulson
parents: 5582
diff changeset
    56
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    57
(*** Monotonicity results ***)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    58
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    59
Goal "(v+z < w+z) = (v < (w::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    60
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    61
qed "zadd_right_cancel_zless";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    62
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    63
Goal "(z+v < z+w) = (v < (w::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    64
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    65
qed "zadd_left_cancel_zless";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    66
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    67
Addsimps [zadd_right_cancel_zless, zadd_left_cancel_zless];
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    68
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    69
Goal "(v+z <= w+z) = (v <= (w::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    70
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    71
qed "zadd_right_cancel_zle";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    72
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    73
Goal "(z+v <= z+w) = (v <= (w::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    74
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    75
qed "zadd_left_cancel_zle";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    76
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    77
Addsimps [zadd_right_cancel_zle, zadd_left_cancel_zle];
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    78
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    79
(*"v<=w ==> v+z <= w+z"*)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    80
bind_thm ("zadd_zless_mono1", zadd_right_cancel_zless RS iffD2);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    81
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    82
(*"v<=w ==> v+z <= w+z"*)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    83
bind_thm ("zadd_zle_mono1", zadd_right_cancel_zle RS iffD2);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    84
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    85
Goal "!!z z'::int. [| w'<=w; z'<=z |] ==> w' + z' <= w + z";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    86
by (etac (zadd_zle_mono1 RS zle_trans) 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    87
by (Simp_tac 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    88
qed "zadd_zle_mono";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    89
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    90
Goal "!!z z'::int. [| w'<w; z'<=z |] ==> w' + z' < w + z";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    91
by (etac (zadd_zless_mono1 RS zless_zle_trans) 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    92
by (Simp_tac 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    93
qed "zadd_zless_mono";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    94
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    95
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    96
(*** Comparison laws ***)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    97
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    98
Goal "(- x < - y) = (y < (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    99
by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   100
qed "zminus_zless_zminus"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   101
Addsimps [zminus_zless_zminus];
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   102
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   103
Goal "(- x <= - y) = (y <= (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   104
by (simp_tac (simpset() addsimps [zle_def]) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   105
qed "zminus_zle_zminus"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   106
Addsimps [zminus_zle_zminus];
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   107
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   108
(** The next several equations can make the simplifier loop! **)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   109
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   110
Goal "(x < - y) = (y < - (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   111
by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   112
qed "zless_zminus"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   113
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   114
Goal "(- x < y) = (- y < (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   115
by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   116
qed "zminus_zless"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   117
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   118
Goal "(x <= - y) = (y <= - (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   119
by (simp_tac (simpset() addsimps [zle_def, zminus_zless]) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   120
qed "zle_zminus"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   121
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   122
Goal "(- x <= y) = (- y <= (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   123
by (simp_tac (simpset() addsimps [zle_def, zless_zminus]) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   124
qed "zminus_zle"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   125
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   126
Goal "- (int (Suc n)) < int 0";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   127
by (simp_tac (simpset() addsimps [zless_def]) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   128
qed "negative_zless_0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   129
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   130
Goal "- (int (Suc n)) < int m";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   131
by (rtac (negative_zless_0 RS zless_zle_trans) 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   132
by (Simp_tac 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   133
qed "negative_zless"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   134
AddIffs [negative_zless]; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   135
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   136
Goal "- int n <= int 0";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   137
by (simp_tac (simpset() addsimps [zminus_zle]) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   138
qed "negative_zle_0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   139
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   140
Goal "- int n <= int m";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   141
by (simp_tac (simpset() addsimps [zless_def, zle_def, zdiff_def, zadd_int]) 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   142
qed "negative_zle"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   143
AddIffs [negative_zle]; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   144
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   145
Goal "~(int 0 <= - (int (Suc n)))";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   146
by (stac zle_zminus 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   147
by (Simp_tac 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   148
qed "not_zle_0_negative"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   149
Addsimps [not_zle_0_negative]; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   150
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   151
Goal "(int n <= - int m) = (n = 0 & m = 0)"; 
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   152
by Safe_tac; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   153
by (Simp_tac 3); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   154
by (dtac (zle_zminus RS iffD1) 2); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   155
by (ALLGOALS (dtac (negative_zle_0 RSN(2,zle_trans)))); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   156
by (ALLGOALS Asm_full_simp_tac); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   157
qed "int_zle_neg"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   158
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   159
Goal "~(int n < - int m)";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   160
by (simp_tac (simpset() addsimps [symmetric zle_def]) 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   161
qed "not_int_zless_negative"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   162
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   163
Goal "(- int n = int m) = (n = 0 & m = 0)"; 
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   164
by (rtac iffI 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   165
by (rtac (int_zle_neg RS iffD1) 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   166
by (dtac sym 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   167
by (ALLGOALS Asm_simp_tac); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   168
qed "negative_eq_positive"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   169
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   170
Addsimps [negative_eq_positive, not_int_zless_negative]; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   171
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   172
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   173
Goal "(w <= z) = (EX n. z = w + int n)";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   174
by (auto_tac (claset() addSIs [not_sym RS not0_implies_Suc],
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   175
	      simpset() addsimps [zless_iff_Suc_zadd, integ_le_less]));
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   176
qed "zle_iff_zadd";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   177
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   178
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   179
Goalw [zdiff_def,zless_def] "neg x = (x < int 0)";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   180
by Auto_tac; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   181
qed "neg_eq_less_nat0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   182
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   183
Goalw [zle_def] "(~neg x) = (int 0 <= x)";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   184
by (simp_tac (simpset() addsimps [neg_eq_less_nat0]) 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   185
qed "not_neg_eq_ge_nat0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   186
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   187
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   188
(**** nat: magnitide of an integer, as a natural number ****)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   189
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   190
Goalw [nat_def] "nat(int n) = n";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   191
by Auto_tac;
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   192
qed "nat_nat";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   193
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   194
Goalw [nat_def] "nat(- int n) = 0";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   195
by (auto_tac (claset(),
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   196
	      simpset() addsimps [neg_eq_less_nat0, zminus_zless])); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   197
qed "nat_zminus_nat";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   198
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   199
Addsimps [nat_nat, nat_zminus_nat];
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   200
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   201
Goal "~ neg z ==> int (nat z) = z"; 
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   202
by (dtac (not_neg_eq_ge_nat0 RS iffD1) 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   203
by (dtac zle_imp_zless_or_eq 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   204
by (auto_tac (claset(), simpset() addsimps [zless_iff_Suc_zadd])); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   205
qed "not_neg_nat"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   206
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   207
Goal "neg x ==> ? n. x = - (int (Suc n))"; 
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   208
by (auto_tac (claset(), 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   209
	      simpset() addsimps [neg_eq_less_nat0, zless_iff_Suc_zadd,
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   210
				  zdiff_eq_eq RS sym, zdiff_def])); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   211
qed "negD"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   212
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   213
Goalw [nat_def] "neg z ==> nat z = 0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   214
by Auto_tac; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   215
qed "neg_nat"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   216
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   217
(* a case theorem distinguishing positive and negative int *)  
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   218
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   219
val prems = Goal "[|!! n. P (int n); !! n. P (- (int (Suc n))) |] ==> P z"; 
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   220
by (case_tac "neg z" 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   221
by (blast_tac (claset() addSDs [negD] addSIs prems) 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   222
by (etac (not_neg_nat RS subst) 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   223
by (resolve_tac prems 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   224
qed "int_cases"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   225
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   226
fun int_case_tac x = res_inst_tac [("z",x)] int_cases; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   227