src/HOL/Integ/Int.ML
author paulson
Tue, 29 Sep 1998 15:57:42 +0200
changeset 5582 a356fb49e69e
parent 5562 02261e6880d1
child 5593 33bca87deae5
permissions -rw-r--r--
many renamings and changes. Simproc for cancelling common terms in relations
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
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    22
(*** Monotonicity results ***)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    23
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    24
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
    25
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    26
qed "zadd_right_cancel_zless";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    27
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    28
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
    29
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    30
qed "zadd_left_cancel_zless";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    31
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    32
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
    33
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    34
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
    35
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    36
qed "zadd_right_cancel_zle";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    37
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    38
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
    39
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    40
qed "zadd_left_cancel_zle";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    41
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    42
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
    43
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    44
(*"v<=w ==> v+z <= w+z"*)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    45
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
    46
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    47
(*"v<=w ==> v+z <= w+z"*)
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    48
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
    49
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    50
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
    51
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
    52
by (Simp_tac 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    53
qed "zadd_zle_mono";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    54
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    55
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
    56
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
    57
by (Simp_tac 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    58
qed "zadd_zless_mono";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    59
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    60
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    61
(*** Comparison laws ***)
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 "(- x < - y) = (y < (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    64
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
    65
qed "zminus_zless_zminus"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    66
Addsimps [zminus_zless_zminus];
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    67
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    68
Goal "(- x <= - y) = (y <= (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    69
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
    70
qed "zminus_zle_zminus"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    71
Addsimps [zminus_zle_zminus];
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
(** 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
    74
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    75
Goal "(x < - y) = (y < - (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    76
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
    77
qed "zless_zminus"; 
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
Goal "(- x < y) = (- y < (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    80
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
    81
qed "zminus_zless"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    82
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    83
Goal "(x <= - y) = (y <= - (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    84
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
    85
qed "zle_zminus"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    86
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    87
Goal "(- x <= y) = (- y <= (x::int))";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    88
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
    89
qed "zminus_zle"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    90
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    91
Goal "- (int (Suc n)) < int 0";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    92
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
    93
qed "negative_zless_0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    94
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
    95
Goal "- (int (Suc n)) < int m";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    96
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
    97
by (Simp_tac 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    98
qed "negative_zless"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
    99
AddIffs [negative_zless]; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   100
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   101
Goal "- int n <= int 0";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   102
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
   103
qed "negative_zle_0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   104
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   105
Goal "- int n <= int m";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   106
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
   107
qed "negative_zle"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   108
AddIffs [negative_zle]; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   109
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   110
Goal "~(int 0 <= - (int (Suc n)))";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   111
by (stac zle_zminus 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   112
by (Simp_tac 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   113
qed "not_zle_0_negative"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   114
Addsimps [not_zle_0_negative]; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   115
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   116
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
   117
by Safe_tac; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   118
by (Simp_tac 3); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   119
by (dtac (zle_zminus RS iffD1) 2); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   120
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
   121
by (ALLGOALS Asm_full_simp_tac); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   122
qed "int_zle_neg"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   123
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   124
Goal "~(int n < - int m)";
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   125
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
   126
qed "not_int_zless_negative"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   127
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   128
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
   129
by (rtac iffI 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   130
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
   131
by (dtac sym 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   132
by (ALLGOALS Asm_simp_tac); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   133
qed "negative_eq_positive"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   134
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   135
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
   136
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   137
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   138
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
   139
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
   140
	      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
   141
qed "zle_iff_zadd";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   142
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   143
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   144
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
   145
by Auto_tac; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   146
qed "neg_eq_less_nat0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   147
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   148
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
   149
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
   150
qed "not_neg_eq_ge_nat0"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   151
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   152
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   153
(**** 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
   154
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   155
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
   156
by Auto_tac;
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   157
qed "nat_nat";
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
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
   160
by (auto_tac (claset(),
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   161
	      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
   162
qed "nat_zminus_nat";
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   163
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   164
Addsimps [nat_nat, nat_zminus_nat];
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   165
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   166
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
   167
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
   168
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
   169
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
   170
qed "not_neg_nat"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   171
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   172
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
   173
by (auto_tac (claset(), 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   174
	      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
   175
				  zdiff_eq_eq RS sym, zdiff_def])); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   176
qed "negD"; 
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
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
   179
by Auto_tac; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   180
qed "neg_nat"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   181
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   182
(* 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
   183
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   184
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
   185
by (case_tac "neg z" 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   186
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
   187
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
   188
by (resolve_tac prems 1);
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   189
qed "int_cases"; 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   190
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents:
diff changeset
   191
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
   192