src/HOL/Integ/Bin.ML
author paulson
Wed, 23 Jun 1999 10:36:59 +0200
changeset 6838 941c4f70db91
parent 6716 87c750df8888
child 6910 7c3503ae3d78
permissions -rw-r--r--
rewrite rules to distribute CONSTANT multiplication over sum and difference; removed automatic rewriting of 2x to x+x
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
     1
(*  Title:      HOL/Integ/Bin.ML
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
     2
    Authors:    Lawrence C Paulson, Cambridge University Computer Laboratory
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
     3
                David Spelt, University of Twente 
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
     4
                Tobias Nipkow, Technical University Munich
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
     5
    Copyright   1994  University of Cambridge
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
     6
    Copyright   1996  University of Twente
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
     7
    Copyright   1999  TU Munich
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
     8
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
     9
Arithmetic on binary integers;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
    10
decision procedure for linear arithmetic.
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    11
*)
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    12
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    13
(** extra rules for bin_succ, bin_pred, bin_add, bin_mult **)
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    14
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    15
qed_goal "NCons_Pls_0" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    16
    "NCons Pls False = Pls"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    17
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    18
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    19
qed_goal "NCons_Pls_1" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    20
    "NCons Pls True = Pls BIT True"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    21
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    22
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    23
qed_goal "NCons_Min_0" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    24
    "NCons Min False = Min BIT False"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    25
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    26
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    27
qed_goal "NCons_Min_1" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    28
    "NCons Min True = Min"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    29
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    30
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    31
qed_goal "bin_succ_1" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    32
    "bin_succ(w BIT True) = (bin_succ w) BIT False"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    33
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    34
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    35
qed_goal "bin_succ_0" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    36
    "bin_succ(w BIT False) =  NCons w True"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    37
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    38
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    39
qed_goal "bin_pred_1" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    40
    "bin_pred(w BIT True) = NCons w False"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    41
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    42
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    43
qed_goal "bin_pred_0" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    44
    "bin_pred(w BIT False) = (bin_pred w) BIT True"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    45
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    46
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    47
qed_goal "bin_minus_1" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    48
    "bin_minus(w BIT True) = bin_pred (NCons (bin_minus w) False)"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    49
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    50
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    51
qed_goal "bin_minus_0" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    52
    "bin_minus(w BIT False) = (bin_minus w) BIT False"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    53
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    54
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
    55
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    56
(*** bin_add: binary addition ***)
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    57
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    58
qed_goal "bin_add_BIT_11" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    59
    "bin_add (v BIT True) (w BIT True) = \
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    60
\    NCons (bin_add v (bin_succ w)) False"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    61
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    62
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    63
qed_goal "bin_add_BIT_10" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    64
    "bin_add (v BIT True) (w BIT False) = NCons (bin_add v w) True"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    65
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    66
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    67
qed_goal "bin_add_BIT_0" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    68
    "bin_add (v BIT False) (w BIT y) = NCons (bin_add v w) y"
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
    69
 (fn _ => [Auto_tac]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    70
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
    71
Goal "bin_add w Pls = w";
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
    72
by (induct_tac "w" 1);
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
    73
by Auto_tac;
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
    74
qed "bin_add_Pls_right";
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    75
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    76
qed_goal "bin_add_BIT_Min" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    77
    "bin_add (v BIT x) Min = bin_pred (v BIT x)"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    78
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    79
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    80
qed_goal "bin_add_BIT_BIT" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    81
    "bin_add (v BIT x) (w BIT y) = \
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    82
\    NCons(bin_add v (if x & y then (bin_succ w) else w)) (x~= y)"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    83
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    84
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    85
6036
1512f4b7d2e8 fixed comments
paulson
parents: 5779
diff changeset
    86
(*** bin_mult: binary multiplication ***)
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    87
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    88
qed_goal "bin_mult_1" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    89
    "bin_mult (v BIT True) w = bin_add (NCons (bin_mult v w) False) w"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    90
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    91
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    92
qed_goal "bin_mult_0" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
    93
    "bin_mult (v BIT False) w = NCons (bin_mult v w) False"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
    94
 (fn _ => [(Simp_tac 1)]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    95
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    96
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    97
(**** The carry/borrow functions, bin_succ and bin_pred ****)
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    98
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
    99
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   100
(**** integ_of ****)
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   101
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   102
qed_goal "integ_of_NCons" Bin.thy
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   103
    "integ_of(NCons w b) = integ_of(w BIT b)"
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5069
diff changeset
   104
 (fn _ =>[(induct_tac "w" 1),
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   105
          (ALLGOALS Asm_simp_tac) ]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   106
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   107
Addsimps [integ_of_NCons];
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   108
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   109
qed_goal "integ_of_succ" Bin.thy
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   110
    "integ_of(bin_succ w) = int 1 + integ_of w"
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
   111
 (fn _ =>[(rtac bin.induct 1),
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   112
          (ALLGOALS(asm_simp_tac (simpset() addsimps zadd_ac))) ]);
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   113
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   114
qed_goal "integ_of_pred" Bin.thy
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   115
    "integ_of(bin_pred w) = - (int 1) + integ_of w"
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   116
 (fn _ =>[(rtac bin.induct 1),
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   117
          (ALLGOALS(asm_simp_tac (simpset() addsimps zadd_ac))) ]);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   118
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   119
Goal "integ_of(bin_minus w) = - (integ_of w)";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   120
by (rtac bin.induct 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   121
by (Simp_tac 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   122
by (Simp_tac 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   123
by (asm_simp_tac (simpset()
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   124
		  delsimps [bin_pred_Pls, bin_pred_Min, bin_pred_BIT]
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   125
		  addsimps [integ_of_succ,integ_of_pred,
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   126
			    zadd_assoc]) 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   127
qed "integ_of_minus";
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   128
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   129
 
6036
1512f4b7d2e8 fixed comments
paulson
parents: 5779
diff changeset
   130
val bin_add_simps = [bin_add_BIT_BIT, integ_of_succ, integ_of_pred];
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   131
6036
1512f4b7d2e8 fixed comments
paulson
parents: 5779
diff changeset
   132
(*This proof is complicated by the mutual recursion*)
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   133
Goal "! w. integ_of(bin_add v w) = integ_of v + integ_of w";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5069
diff changeset
   134
by (induct_tac "v" 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4641
diff changeset
   135
by (simp_tac (simpset() addsimps bin_add_simps) 1);
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4641
diff changeset
   136
by (simp_tac (simpset() addsimps bin_add_simps) 1);
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   137
by (rtac allI 1);
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5069
diff changeset
   138
by (induct_tac "w" 1);
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   139
by (ALLGOALS (asm_simp_tac (simpset() addsimps bin_add_simps @ zadd_ac)));
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   140
qed_spec_mp "integ_of_add";
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   141
5779
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   142
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   143
(*Subtraction*)
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   144
Goalw [zdiff_def]
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   145
     "integ_of v - integ_of w = integ_of(bin_add v (bin_minus w))";
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   146
by (simp_tac (simpset() addsimps [integ_of_add, integ_of_minus]) 1);
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   147
qed "diff_integ_of_eq";
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   148
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   149
val bin_mult_simps = [zmult_zminus, integ_of_minus, integ_of_add];
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   150
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   151
Goal "integ_of(bin_mult v w) = integ_of v * integ_of w";
5184
9b8547a9496a Adapted to new datatype package.
berghofe
parents: 5069
diff changeset
   152
by (induct_tac "v" 1);
4686
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4641
diff changeset
   153
by (simp_tac (simpset() addsimps bin_mult_simps) 1);
74a12e86b20b Removed `addsplits [expand_if]'
nipkow
parents: 4641
diff changeset
   154
by (simp_tac (simpset() addsimps bin_mult_simps) 1);
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   155
by (asm_simp_tac
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   156
    (simpset() addsimps bin_mult_simps @ [zadd_zmult_distrib] @ zadd_ac) 1);
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   157
qed "integ_of_mult";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   158
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   159
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   160
(** Simplification rules with integer constants **)
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   161
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   162
Goal "#0 + z = z";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   163
by (Simp_tac 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   164
qed "zadd_0";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   165
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   166
Goal "z + #0 = z";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   167
by (Simp_tac 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   168
qed "zadd_0_right";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   169
5592
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   170
Addsimps [zadd_0, zadd_0_right];
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   171
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   172
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   173
(** Converting simple cases of (int n) to numerals **)
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   174
5592
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   175
(*int 0 = #0 *)
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   176
bind_thm ("int_0", integ_of_Pls RS sym);
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   177
5592
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   178
Goal "int (Suc n) = #1 + int n";
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   179
by (simp_tac (simpset() addsimps [zadd_int]) 1);
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   180
qed "int_Suc";
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   181
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   182
Goal "- (#0) = #0";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   183
by (Simp_tac 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   184
qed "zminus_0";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   185
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   186
Addsimps [zminus_0];
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   187
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   188
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   189
Goal "#0 - x = -x";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   190
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   191
qed "zdiff0";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   192
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   193
Goal "x - #0 = x";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   194
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   195
qed "zdiff0_right";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   196
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   197
Goal "x - x = #0";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   198
by (simp_tac (simpset() addsimps [zdiff_def]) 1);
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   199
qed "zdiff_self";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   200
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   201
Addsimps [zdiff0, zdiff0_right, zdiff_self];
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   202
6838
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   203
(** Distributive laws for constants only **)
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   204
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   205
Addsimps (map (read_instantiate_sg (sign_of Bin.thy) [("w", "integ_of ?v")])
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   206
	  [zadd_zmult_distrib, zadd_zmult_distrib2,
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   207
	   zdiff_zmult_distrib, zdiff_zmult_distrib2]);
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   208
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   209
(** Special-case simplification for small constants **)
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   210
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   211
Goal "#0 * z = #0";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   212
by (Simp_tac 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   213
qed "zmult_0";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   214
6838
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   215
Goal "z * #0 = #0";
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   216
by (Simp_tac 1);
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   217
qed "zmult_0_right";
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   218
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   219
Goal "#1 * z = z";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   220
by (Simp_tac 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   221
qed "zmult_1";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   222
6838
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   223
Goal "z * #1 = z";
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   224
by (Simp_tac 1);
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   225
qed "zmult_1_right";
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   226
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   227
(*For specialist use*)
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   228
Goal "#2 * z = z+z";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   229
by (simp_tac (simpset() addsimps [zadd_zmult_distrib]) 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   230
qed "zmult_2";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   231
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   232
Goal "z * #2 = z+z";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   233
by (simp_tac (simpset() addsimps [zadd_zmult_distrib2]) 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   234
qed "zmult_2_right";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   235
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   236
Addsimps [zmult_0, zmult_0_right, 
6838
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   237
	  zmult_1, zmult_1_right];
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   238
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   239
Goal "(w < z + #1) = (w<z | w=z)";
5592
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   240
by (simp_tac (simpset() addsimps [zless_add_int_Suc_eq]) 1);
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   241
qed "zless_add1_eq";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   242
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   243
Goal "(w + #1 <= z) = (w<z)";
5592
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   244
by (simp_tac (simpset() addsimps [add_int_Suc_zle_eq]) 1);
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   245
qed "add1_zle_eq";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   246
Addsimps [add1_zle_eq];
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   247
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   248
Goal "neg x = (x < #0)";
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   249
by (simp_tac (simpset() addsimps [neg_eq_less_nat0]) 1); 
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   250
qed "neg_eq_less_0"; 
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   251
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   252
Goal "(~neg x) = (int 0 <= x)";
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   253
by (simp_tac (simpset() addsimps [not_neg_eq_ge_nat0]) 1); 
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   254
qed "not_neg_eq_ge_0"; 
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   255
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   256
Goal "#0 <= int m";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   257
by (Simp_tac 1);
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   258
qed "zero_zle_int";
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   259
AddIffs [zero_zle_int];
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   260
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   261
5747
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   262
(** Needed because (int 0) rewrites to #0.
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   263
    Can these be generalized without evaluating large numbers?**)
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   264
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   265
Goal "~ (int k < #0)";
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   266
by (Simp_tac 1);
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   267
qed "int_less_0_conv";
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   268
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   269
Goal "(int k <= #0) = (k=0)";
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   270
by (Simp_tac 1);
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   271
qed "int_le_0_conv";
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   272
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   273
Goal "(int k = #0) = (k=0)";
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   274
by (Simp_tac 1);
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   275
qed "int_eq_0_conv";
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   276
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   277
Goal "(#0 = int k) = (k=0)";
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   278
by Auto_tac;
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   279
qed "int_eq_0_conv'";
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   280
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   281
Addsimps [int_less_0_conv, int_le_0_conv, int_eq_0_conv, int_eq_0_conv'];
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   282
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   283
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   284
(** Simplification rules for comparison of binary numbers (Norbert Voelker) **)
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   285
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   286
(** Equals (=) **)
1632
39e146ac224c Binary integers and their numeric syntax
paulson
parents:
diff changeset
   287
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   288
Goalw [iszero_def]
5779
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   289
  "(integ_of x = integ_of y) = iszero(integ_of (bin_add x (bin_minus y)))"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   290
by (simp_tac (simpset() addsimps
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   291
              (zcompare_rls @ [integ_of_add, integ_of_minus])) 1); 
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   292
qed "eq_integ_of_eq"; 
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   293
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   294
Goalw [iszero_def] "iszero (integ_of Pls)"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   295
by (Simp_tac 1); 
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   296
qed "iszero_integ_of_Pls"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   297
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   298
Goalw [iszero_def] "~ iszero(integ_of Min)"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   299
by (Simp_tac 1);
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   300
qed "nonzero_integ_of_Min"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   301
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   302
Goalw [iszero_def]
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   303
     "iszero (integ_of (w BIT x)) = (~x & iszero (integ_of w))"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   304
by (Simp_tac 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   305
by (int_case_tac "integ_of w" 1); 
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   306
by (ALLGOALS (asm_simp_tac 
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   307
	      (simpset() addsimps zcompare_rls @ 
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   308
				  [zminus_zadd_distrib RS sym, 
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   309
				   zadd_int]))); 
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   310
qed "iszero_integ_of_BIT"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   311
5779
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   312
Goal "iszero (integ_of (w BIT False)) = iszero (integ_of w)"; 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   313
by (simp_tac (HOL_ss addsimps [iszero_integ_of_BIT]) 1); 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   314
qed "iszero_integ_of_0"; 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   315
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   316
Goal "~ iszero (integ_of (w BIT True))"; 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   317
by (simp_tac (HOL_ss addsimps [iszero_integ_of_BIT]) 1); 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   318
qed "iszero_integ_of_1"; 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   319
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   320
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   321
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   322
(** Less-than (<) **)
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   323
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   324
Goalw [zless_def,zdiff_def] 
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   325
    "integ_of x < integ_of y \
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   326
\    = neg (integ_of (bin_add x (bin_minus y)))";
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   327
by (simp_tac (simpset() addsimps bin_mult_simps) 1);
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   328
qed "less_integ_of_eq_neg"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   329
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   330
Goal "~ neg (integ_of Pls)"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   331
by (Simp_tac 1); 
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   332
qed "not_neg_integ_of_Pls"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   333
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   334
Goal "neg (integ_of Min)"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   335
by (Simp_tac 1);
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   336
qed "neg_integ_of_Min"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   337
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   338
Goal "neg (integ_of (w BIT x)) = neg (integ_of w)"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   339
by (Asm_simp_tac 1); 
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   340
by (int_case_tac "integ_of w" 1); 
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   341
by (ALLGOALS (asm_simp_tac 
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   342
	      (simpset() addsimps [zadd_int, neg_eq_less_nat0, 
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   343
				   symmetric zdiff_def] @ zcompare_rls))); 
5512
4327eec06849 much renaming and tidying
paulson
parents: 5510
diff changeset
   344
qed "neg_integ_of_BIT"; 
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   345
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   346
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   347
(** Less-than-or-equals (<=) **)
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   348
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   349
Goal "(integ_of x <= integ_of y) = (~ integ_of y < integ_of x)";
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   350
by (simp_tac (simpset() addsimps [zle_def]) 1);
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   351
qed "le_integ_of_eq_not_less"; 
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   352
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   353
(*Delete the original rewrites, with their clumsy conditional expressions*)
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   354
Delsimps [bin_succ_BIT, bin_pred_BIT, bin_minus_BIT, 
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   355
          NCons_Pls, NCons_Min, bin_add_BIT, bin_mult_BIT];
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   356
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   357
(*Hide the binary representation of integer constants*)
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   358
Delsimps [integ_of_Pls, integ_of_Min, integ_of_BIT];
5491
22f8331cdf47 revised treatment of integers
paulson
parents: 5224
diff changeset
   359
5779
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   360
(*simplification of arithmetic operations on integer constants*)
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   361
val bin_arith_extra_simps =
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   362
    [integ_of_add RS sym,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   363
     integ_of_minus RS sym,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   364
     integ_of_mult RS sym,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   365
     bin_succ_1, bin_succ_0, 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   366
     bin_pred_1, bin_pred_0, 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   367
     bin_minus_1, bin_minus_0,  
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   368
     bin_add_Pls_right, bin_add_BIT_Min,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   369
     bin_add_BIT_0, bin_add_BIT_10, bin_add_BIT_11,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   370
     diff_integ_of_eq, 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   371
     bin_mult_1, bin_mult_0, 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   372
     NCons_Pls_0, NCons_Pls_1, 
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   373
     NCons_Min_0, NCons_Min_1, NCons_BIT];
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
   374
5779
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   375
(*For making a minimal simpset, one must include these default simprules
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   376
  of Bin.thy.  Also include simp_thms, or at least (~False)=True*)
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   377
val bin_arith_simps =
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   378
    [bin_pred_Pls, bin_pred_Min,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   379
     bin_succ_Pls, bin_succ_Min,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   380
     bin_add_Pls, bin_add_Min,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   381
     bin_minus_Pls, bin_minus_Min,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   382
     bin_mult_Pls, bin_mult_Min] @ bin_arith_extra_simps;
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
   383
5779
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   384
(*Simplification of relational operations*)
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   385
val bin_rel_simps =
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   386
    [eq_integ_of_eq, iszero_integ_of_Pls, nonzero_integ_of_Min,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   387
     iszero_integ_of_0, iszero_integ_of_1,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   388
     less_integ_of_eq_neg,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   389
     not_neg_integ_of_Pls, neg_integ_of_Min, neg_integ_of_BIT,
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   390
     le_integ_of_eq_not_less];
2224
4fc4b465be5b New material from Norbert Voelker for efficient binary comparisons
paulson
parents: 1894
diff changeset
   391
5779
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   392
Addsimps bin_arith_extra_simps;
5c74f003a68e Explicit (and improved) simprules for binary arithmetic.
paulson
parents: 5747
diff changeset
   393
Addsimps bin_rel_simps;
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   394
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   395
(*---------------------------------------------------------------------------*)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   396
(* Linear arithmetic                                                         *)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   397
(*---------------------------------------------------------------------------*)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   398
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   399
(*
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   400
Instantiation of the generic linear arithmetic package for int.
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   401
FIXME: multiplication with constants (eg #2 * i) does not work yet.
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   402
Solution: the cancellation simprocs in Int_Cancel should be able to deal with
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   403
it (eg simplify #3 * i <= 2 * i to i <= #0) or `add_rules' below should
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   404
include rules for turning multiplication with constants into addition.
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   405
(The latter option is very inefficient!)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   406
*)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   407
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   408
structure Int_LA_Data(*: LIN_ARITH_DATA*) =
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   409
struct
6101
dde00dc06f0d Restructured Arithmatic
nipkow
parents: 6079
diff changeset
   410
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   411
val lessD = Nat_LA_Data.lessD @ [add1_zle_eq RS iffD2];
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   412
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   413
(* borrowed from Bin.thy: *)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   414
fun dest_bit (Const ("False", _)) = 0
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   415
  | dest_bit (Const ("True", _)) = 1
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   416
  | dest_bit _ = raise Match;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   417
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   418
fun dest_bin tm =
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   419
  let
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   420
    fun bin_of (Const ("Bin.bin.Pls", _)) = []
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   421
      | bin_of (Const ("Bin.bin.Min", _)) = [~1]
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   422
      | bin_of (Const ("Bin.bin.op BIT", _) $ bs $ b) = dest_bit b :: bin_of bs
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   423
      | bin_of _ = raise Match;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   424
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   425
    fun int_of [] = 0
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   426
      | int_of (b :: bs) = b + 2 * int_of bs;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   427
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   428
  in int_of(bin_of tm) end;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   429
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   430
fun add_atom(t,m,(p,i)) = (case assoc(p,t) of None => ((t,m)::p,i)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   431
                           | Some n => (overwrite(p,(t,n+m:int)), i));
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   432
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   433
(* Turn term into list of summand * multiplicity plus a constant *)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   434
fun poly(Const("op +",_) $ s $ t, m, pi) = poly(s,m,poly(t,m,pi))
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   435
  | poly(Const("op -",_) $ s $ t, m, pi) = poly(s,m,poly(t,~1*m,pi))
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   436
  | poly(Const("uminus",_) $ t, m, pi) =   poly(t,~1*m,pi)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   437
  | poly(all as Const("op *",_) $ (Const("Bin.integ_of",_)$c) $ t, m, pi) =
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   438
      (poly(t,m*dest_bin c,pi) handle Match => add_atom(all,m,pi))
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   439
  | poly(all as Const("Bin.integ_of",_)$t,m,(p,i)) =
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   440
     ((p,i + m*dest_bin t) handle Match => add_atom(all,m,(p,i)))
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   441
  | poly x  = add_atom x;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   442
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   443
fun decomp2(rel,lhs,rhs) =
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   444
  let val (p,i) = poly(lhs,1,([],0)) and (q,j) = poly(rhs,1,([],0))
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   445
  in case rel of
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   446
       "op <"  => Some(p,i,"<",q,j)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   447
     | "op <=" => Some(p,i,"<=",q,j)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   448
     | "op ="  => Some(p,i,"=",q,j)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   449
     | _       => None
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   450
  end;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   451
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   452
val intT = Type("IntDef.int",[]);
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   453
fun iib T = T = ([intT,intT] ---> HOLogic.boolT);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   454
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   455
fun decomp1(T,xxx) =
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   456
  if iib T then decomp2 xxx else Nat_LA_Data.decomp1(T,xxx);
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   457
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   458
fun decomp(_$(Const(rel,T)$lhs$rhs)) = decomp1(T,(rel,lhs,rhs))
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   459
  | decomp(_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) =
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   460
      Nat_LA_Data.negate(decomp1(T,(rel,lhs,rhs)))
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   461
  | decomp _ = None
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   462
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   463
(* reduce contradictory <= to False *)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   464
val add_rules = simp_thms@bin_arith_simps@bin_rel_simps@[int_0];
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   465
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   466
val cancel_sums_ss = Nat_LA_Data.cancel_sums_ss addsimps add_rules
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   467
          addsimprocs [Int_Cancel.sum_conv, Int_Cancel.rel_conv];
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   468
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   469
val simp = simplify cancel_sums_ss;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   470
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   471
val add_mono_thms = Nat_LA_Data.add_mono_thms @
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   472
  map (fn s => prove_goal Int.thy s
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   473
                 (fn prems => [cut_facts_tac prems 1,
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   474
                      asm_simp_tac (simpset() addsimps [zadd_zle_mono]) 1]))
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   475
    ["(i <= j) & (k <= l) ==> i + k <= j + (l::int)",
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   476
     "(i  = j) & (k <= l) ==> i + k <= j + (l::int)",
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   477
     "(i <= j) & (k  = l) ==> i + k <= j + (l::int)",
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   478
     "(i  = j) & (k  = l) ==> i + k  = j + (l::int)"
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   479
    ];
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   480
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   481
end;
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   482
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   483
(* Update parameters of arithmetic prover *)
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   484
LA_Data_Ref.add_mono_thms := Int_LA_Data.add_mono_thms;
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   485
LA_Data_Ref.lessD :=         Int_LA_Data.lessD;
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   486
LA_Data_Ref.decomp :=        Int_LA_Data.decomp;
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   487
LA_Data_Ref.simp :=          Int_LA_Data.simp;
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   488
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   489
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   490
val int_arith_simproc_pats =
6394
3d9fd50fcc43 Theory.sign_of;
wenzelm
parents: 6301
diff changeset
   491
  map (fn s => Thm.read_cterm (Theory.sign_of Int.thy) (s, HOLogic.boolT))
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   492
      ["(m::int) < n","(m::int) <= n", "(m::int) = n"];
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   493
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   494
val fast_int_arith_simproc = mk_simproc "fast_int_arith" int_arith_simproc_pats
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   495
                                        Fast_Arith.lin_arith_prover;
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   496
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   497
Addsimprocs [fast_int_arith_simproc];
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   498
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   499
(* FIXME: K true should be replaced by a sensible test to speed things up
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   500
   in case there are lots of irrelevant terms involved.
6157
29942d3a1818 arith_tac for min/max
nipkow
parents: 6128
diff changeset
   501
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   502
val arith_tac =
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   503
  refute_tac (K true) (REPEAT o split_tac[nat_diff_split])
2acc5d36610c More arith refinements.
nipkow
parents: 6109
diff changeset
   504
             ((REPEAT_DETERM o etac linorder_neqE) THEN' fast_arith_tac);
6157
29942d3a1818 arith_tac for min/max
nipkow
parents: 6128
diff changeset
   505
*)
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   506
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   507
(* Some test data
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   508
Goal "!!a::int. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   509
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   510
Goal "!!a::int. [| a < b; c < d |] ==> a-d+ #2 <= b+(-c)";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   511
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   512
Goal "!!a::int. [| a < b; c < d |] ==> a+c+ #1 < b+d";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   513
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   514
Goal "!!a::int. [| a <= b; b+b <= c |] ==> a+a <= c";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   515
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   516
Goal "!!a::int. [| a+b <= i+j; a<=b; i<=j |] \
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   517
\     ==> a+a <= j+j";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   518
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   519
Goal "!!a::int. [| a+b < i+j; a<b; i<j |] \
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   520
\     ==> a+a - - #-1 < j+j - #3";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   521
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   522
Goal "!!a::int. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   523
by (arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   524
Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   525
\     ==> a <= l";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   526
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   527
Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   528
\     ==> a+a+a+a <= l+l+l+l";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   529
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   530
Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   531
\     ==> a+a+a+a+a <= l+l+l+l+i";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   532
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   533
Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   534
\     ==> a+a+a+a+a+a <= l+l+l+l+i+l";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   535
by (fast_arith_tac 1);
6060
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   536
*)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   537
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   538
(*---------------------------------------------------------------------------*)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   539
(* End of linear arithmetic                                                  *)
d30d1dd2082d Instantiated lin.arith.
nipkow
parents: 6036
diff changeset
   540
(*---------------------------------------------------------------------------*)
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   541
5592
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   542
(** Simplification of arithmetic when nested to the right **)
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   543
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   544
Goal "integ_of v + (integ_of w + z) = integ_of(bin_add v w) + z";
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   545
by (simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1);
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   546
qed "add_integ_of_left";
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   547
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   548
Goal "integ_of v * (integ_of w * z) = integ_of(bin_mult v w) * z";
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   549
by (simp_tac (simpset() addsimps [zmult_assoc RS sym]) 1);
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   550
qed "mult_integ_of_left";
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   551
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   552
Addsimps [add_integ_of_left, mult_integ_of_left];
64697e426048 better handling of literals
paulson
parents: 5582
diff changeset
   553
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   554
(** Simplification of inequalities involving numerical constants **)
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   555
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   556
Goal "(w <= z + #1) = (w<=z | w = z + #1)";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   557
by (arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   558
qed "zle_add1_eq";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   559
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   560
Goal "(w <= z - #1) = (w<z)";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   561
by (arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   562
qed "zle_diff1_eq";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   563
Addsimps [zle_diff1_eq];
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   564
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   565
(*2nd premise can be proved automatically if v is a literal*)
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   566
Goal "[| w <= z; #0 <= v |] ==> w <= z + v";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   567
by (fast_arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   568
qed "zle_imp_zle_zadd";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   569
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   570
Goal "w <= z ==> w <= z + #1";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   571
by (fast_arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   572
qed "zle_imp_zle_zadd1";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   573
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   574
(*2nd premise can be proved automatically if v is a literal*)
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   575
Goal "[| w < z; #0 <= v |] ==> w < z + v";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   576
by (fast_arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   577
qed "zless_imp_zless_zadd";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   578
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   579
Goal "w < z ==> w < z + #1";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   580
by (fast_arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   581
qed "zless_imp_zless_zadd1";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   582
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   583
Goal "(w < z + #1) = (w<=z)";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   584
by (arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   585
qed "zle_add1_eq_le";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   586
Addsimps [zle_add1_eq_le];
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   587
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   588
Goal "(z = z + w) = (w = #0)";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   589
by (arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   590
qed "zadd_left_cancel0";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   591
Addsimps [zadd_left_cancel0];
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   592
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   593
(*LOOPS as a simprule!*)
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   594
Goal "[| w + v < z; #0 <= v |] ==> w < z";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   595
by (fast_arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   596
qed "zless_zadd_imp_zless";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   597
5540
0f16c3b66ab4 much renaming and reorganization
paulson
parents: 5512
diff changeset
   598
(*LOOPS as a simprule!  Analogous to Suc_lessD*)
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   599
Goal "w + #1 < z ==> w < z";
6301
08245f5a436d expandshort
paulson
parents: 6157
diff changeset
   600
by (fast_arith_tac 1);
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   601
qed "zless_zadd1_imp_zless";
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   602
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   603
Goal "w + #-1 = w - #1";
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   604
by (Simp_tac 1);
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   605
qed "zplus_minus1_conv";
5510
ad120f7c52ad improved (but still flawed) treatment of binary arithmetic
paulson
parents: 5491
diff changeset
   606
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   607
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   608
(*** nat ***)
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   609
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   610
Goal "#0 <= z ==> int (nat z) = z"; 
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   611
by (asm_full_simp_tac
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   612
    (simpset() addsimps [neg_eq_less_0, zle_def, not_neg_nat]) 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   613
qed "nat_0_le"; 
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   614
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   615
Goal "z < #0 ==> nat z = 0"; 
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   616
by (asm_full_simp_tac
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   617
    (simpset() addsimps [neg_eq_less_0, zle_def, neg_nat]) 1); 
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   618
qed "nat_less_0"; 
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   619
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   620
Addsimps [nat_0_le, nat_less_0];
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   621
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   622
Goal "#0 <= w ==> (nat w = m) = (w = int m)";
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   623
by Auto_tac;
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   624
qed "nat_eq_iff";
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   625
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   626
Goal "#0 <= w ==> (nat w < m) = (w < int m)";
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   627
by (rtac iffI 1);
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   628
by (asm_full_simp_tac 
5582
a356fb49e69e many renamings and changes. Simproc for cancelling common terms in relations
paulson
parents: 5562
diff changeset
   629
    (simpset() delsimps [zless_int] addsimps [zless_int RS sym]) 2);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   630
by (etac (nat_0_le RS subst) 1);
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   631
by (Simp_tac 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   632
qed "nat_less_iff";
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   633
5747
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   634
6716
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   635
(*Users don't want to see (int 0), int(Suc 0) or w + - z*)
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   636
Addsimps [int_0, int_Suc, symmetric zdiff_def];
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   637
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   638
Goal "nat #0 = 0";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   639
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   640
qed "nat_0";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   641
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   642
Goal "nat #1 = 1";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   643
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   644
qed "nat_1";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   645
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   646
Goal "nat #2 = 2";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   647
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   648
qed "nat_2";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   649
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   650
Goal "nat #3 = Suc 2";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   651
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   652
qed "nat_3";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   653
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   654
Goal "nat #4 = Suc (Suc 2)";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   655
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   656
qed "nat_4";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   657
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   658
Goal "nat #5 = Suc (Suc (Suc 2))";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   659
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   660
qed "nat_5";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   661
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   662
Goal "nat #6 = Suc (Suc (Suc (Suc 2)))";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   663
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   664
qed "nat_6";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   665
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   666
Goal "nat #7 = Suc (Suc (Suc (Suc (Suc 2))))";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   667
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   668
qed "nat_7";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   669
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   670
Goal "nat #8 = Suc (Suc (Suc (Suc (Suc (Suc 2)))))";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   671
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   672
qed "nat_8";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   673
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   674
Goal "nat #9 = Suc (Suc (Suc (Suc (Suc (Suc (Suc 2))))))";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   675
by (simp_tac (simpset() addsimps [nat_eq_iff]) 1);
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   676
qed "nat_9";
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   677
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   678
(*Users also don't want to see (nat 0), (nat 1), ...*)
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   679
Addsimps [nat_0, nat_1, nat_2, nat_3, nat_4, 
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   680
	  nat_5, nat_6, nat_7, nat_8, nat_9];
87c750df8888 Better simplification of (nat #0), (int (Suc 0)), etc
paulson
parents: 6394
diff changeset
   681
5747
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   682
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   683
Goal "#0 <= w ==> (nat w < nat z) = (w<z)";
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   684
by (case_tac "neg z" 1);
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   685
by (auto_tac (claset(), simpset() addsimps [nat_less_iff]));
5551
ed5e19bc7e32 renamed some axioms; some new theorems
paulson
parents: 5540
diff changeset
   686
by (auto_tac (claset() addIs [zless_trans], 
5747
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   687
	      simpset() addsimps [neg_eq_less_0, zle_def]));
5562
02261e6880d1 Renaming of Integ/Integ.* to Integ/Int.*, and renaming of related constants
paulson
parents: 5551
diff changeset
   688
qed "nat_less_eq_zless";
5747
387b5bf9326a Now users will never see (int 0)
paulson
parents: 5592
diff changeset
   689
6838
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   690
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   691
Addsimps zadd_ac;
941c4f70db91 rewrite rules to distribute CONSTANT multiplication over sum and difference;
paulson
parents: 6716
diff changeset
   692
Addsimps zmult_ac;