src/HOL/Integ/nat_bin.ML
author nipkow
Mon, 01 Jul 2002 12:50:35 +0200
changeset 13261 a0460a450cf9
parent 13183 c7290200b3f4
child 13462 56610e2ba220
permissions -rw-r--r--
fixed problem with linear arith.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/nat_bin.ML
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     5
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     6
Binary arithmetic for the natural numbers
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     7
*)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     8
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
     9
val nat_number_of_def = thm "nat_number_of_def";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    10
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    11
(** nat (coercion from int to nat) **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    12
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    13
Goal "nat (number_of w) = number_of w";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    14
by (simp_tac (simpset() addsimps [nat_number_of_def]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    15
qed "nat_number_of";
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    16
Addsimps [nat_number_of, nat_0, nat_1];
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    17
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11464
diff changeset
    18
Goal "Numeral0 = (0::nat)";
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    19
by (simp_tac (simpset() addsimps [nat_number_of_def]) 1); 
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    20
qed "numeral_0_eq_0";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    21
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11464
diff changeset
    22
Goal "Numeral1 = (1::nat)";
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    23
by (simp_tac (simpset() addsimps [nat_1, nat_number_of_def]) 1); 
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    24
qed "numeral_1_eq_1";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    25
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    26
Goal "Numeral1 = Suc 0";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    27
by (simp_tac (simpset() addsimps [numeral_1_eq_1]) 1); 
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    28
qed "numeral_1_eq_Suc_0";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    29
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 11868
diff changeset
    30
Goalw [nat_number_of_def] "2 = Suc (Suc 0)";
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    31
by (rtac nat_2 1); 
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    32
qed "numeral_2_eq_2";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    33
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    34
(** int (coercion from nat to int) **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    35
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    36
(*"neg" is used in rewrite rules for binary comparisons*)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    37
Goal "int (number_of v :: nat) = \
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    38
\        (if neg (number_of v) then 0 \
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    39
\         else (number_of v :: int))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    40
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    41
    (simpset_of Int.thy addsimps [neg_nat, nat_number_of_def, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    42
				  not_neg_nat, int_0]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    43
qed "int_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    44
Addsimps [int_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    45
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    46
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    47
(** Successor **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    48
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    49
Goal "(0::int) <= z ==> Suc (nat z) = nat (1 + z)";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    50
by (rtac sym 1);
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    51
by (asm_simp_tac (simpset() addsimps [nat_eq_iff, int_Suc]) 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    52
qed "Suc_nat_eq_nat_zadd1";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    53
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    54
Goal "Suc (number_of v + n) = \
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    55
\       (if neg (number_of v) then 1+n else number_of (bin_succ v) + n)";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    56
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    57
    (simpset_of Int.thy addsimps [neg_nat, nat_1, not_neg_eq_ge_0, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    58
				  nat_number_of_def, int_Suc, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    59
				  Suc_nat_eq_nat_zadd1, number_of_succ]) 1);
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    60
qed "Suc_nat_number_of_add";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    61
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    62
Goal "Suc (number_of v) = \
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    63
\       (if neg (number_of v) then 1 else number_of (bin_succ v))";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    64
by (cut_inst_tac [("n","0")] Suc_nat_number_of_add 1);
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    65
by (asm_full_simp_tac (simpset() delcongs [if_weak_cong]) 1); 
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    66
qed "Suc_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    67
Addsimps [Suc_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    68
13261
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    69
val nat_bin_arith_setup =
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    70
 [Fast_Arith.map_data 
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    71
   (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} =>
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    72
     {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    73
      inj_thms = inj_thms,
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    74
      lessD = lessD,
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    75
      simpset = simpset addsimps [Suc_nat_number_of, int_nat_number_of,
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    76
                                  not_neg_number_of_Pls,
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    77
                                  neg_number_of_Min,neg_number_of_BIT]})];
a0460a450cf9 fixed problem with linear arith.
nipkow
parents: 13183
diff changeset
    78
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    79
(** Addition **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    80
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
    81
Goal "[| (0::int) <= z;  0 <= z' |] ==> nat (z+z') = nat z + nat z'";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    82
by (rtac (inj_int RS injD) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    83
by (asm_simp_tac (simpset() addsimps [zadd_int RS sym]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    84
qed "nat_add_distrib";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    85
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    86
(*"neg" is used in rewrite rules for binary comparisons*)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    87
Goal "(number_of v :: nat) + number_of v' = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    88
\        (if neg (number_of v) then number_of v' \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    89
\         else if neg (number_of v') then number_of v \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    90
\         else number_of (bin_add v v'))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    91
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    92
    (simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    93
				  nat_add_distrib RS sym, number_of_add]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    94
qed "add_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    95
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    96
Addsimps [add_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    97
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    98
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
    99
(** Subtraction **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   100
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   101
Goal "[| (0::int) <= z';  z' <= z |] ==> nat (z-z') = nat z - nat z'";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   102
by (rtac (inj_int RS injD) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   103
by (asm_simp_tac (simpset() addsimps [zdiff_int RS sym, nat_le_eq_zle]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   104
qed "nat_diff_distrib";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   105
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   106
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   107
Goal "nat z - nat z' = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   108
\       (if neg z' then nat z  \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   109
\        else let d = z-z' in    \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   110
\             if neg d then 0 else nat d)";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   111
by (simp_tac (simpset() addsimps [Let_def, nat_diff_distrib RS sym,
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   112
				  neg_eq_less_0, not_neg_eq_ge_0]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   113
by (simp_tac (simpset() addsimps [diff_is_0_eq, nat_le_eq_zle]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   114
qed "diff_nat_eq_if";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   115
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   116
Goalw [nat_number_of_def]
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   117
     "(number_of v :: nat) - number_of v' = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   118
\       (if neg (number_of v') then number_of v \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   119
\        else let d = number_of (bin_add v (bin_minus v')) in    \
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   120
\             if neg d then 0 else nat d)";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   121
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   122
    (simpset_of Int.thy delcongs [if_weak_cong]
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   123
			addsimps [not_neg_eq_ge_0, nat_0,
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   124
				  diff_nat_eq_if, diff_number_of_eq]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   125
qed "diff_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   126
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   127
Addsimps [diff_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   128
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   129
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   130
(** Multiplication **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   131
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   132
Goal "(0::int) <= z ==> nat (z*z') = nat z * nat z'";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   133
by (case_tac "0 <= z'" 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   134
by (asm_full_simp_tac (simpset() addsimps [zmult_le_0_iff]) 2);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   135
by (rtac (inj_int RS injD) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   136
by (asm_simp_tac (simpset() addsimps [zmult_int RS sym,
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   137
				      int_0_le_mult_iff]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   138
qed "nat_mult_distrib";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   139
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   140
Goal "z <= (0::int) ==> nat(z*z') = nat(-z) * nat(-z')"; 
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   141
by (rtac trans 1); 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   142
by (rtac nat_mult_distrib 2); 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   143
by Auto_tac;  
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   144
qed "nat_mult_distrib_neg";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   145
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   146
Goal "(number_of v :: nat) * number_of v' = \
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   147
\      (if neg (number_of v) then 0 else number_of (bin_mult v v'))";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   148
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   149
    (simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   150
				  nat_mult_distrib RS sym, number_of_mult, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   151
				  nat_0]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   152
qed "mult_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   153
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   154
Addsimps [mult_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   155
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   156
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   157
(** Quotient **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   158
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   159
Goal "(0::int) <= z ==> nat (z div z') = nat z div nat z'";
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   160
by (case_tac "0 <= z'" 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   161
by (auto_tac (claset(), 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   162
	      simpset() addsimps [div_nonneg_neg_le0, DIVISION_BY_ZERO_DIV]));
13183
c7290200b3f4 conversion of IntDiv.thy to Isar format
paulson
parents: 12613
diff changeset
   163
by (case_tac "z' = 0" 1 THEN asm_simp_tac (simpset() addsimps [DIVISION_BY_ZERO]) 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   164
 by (simp_tac (simpset() addsimps [numeral_0_eq_0, DIVISION_BY_ZERO_DIV]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   165
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset()));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   166
by (rename_tac "m m'" 1);
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   167
by (subgoal_tac "0 <= int m div int m'" 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   168
 by (asm_full_simp_tac 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   169
     (simpset() addsimps [numeral_0_eq_0, pos_imp_zdiv_nonneg_iff]) 2);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   170
by (rtac (inj_int RS injD) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   171
by (Asm_simp_tac 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   172
by (res_inst_tac [("r", "int (m mod m')")] quorem_div 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   173
 by (Force_tac 2);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   174
by (asm_full_simp_tac 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   175
    (simpset() addsimps [nat_less_iff RS sym, quorem_def, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   176
	                 numeral_0_eq_0, zadd_int, zmult_int]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   177
by (rtac (mod_div_equality RS sym RS trans) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   178
by (asm_simp_tac (simpset() addsimps add_ac@mult_ac) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   179
qed "nat_div_distrib";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   180
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   181
Goal "(number_of v :: nat)  div  number_of v' = \
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   182
\         (if neg (number_of v) then 0 \
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   183
\          else nat (number_of v div number_of v'))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   184
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   185
    (simpset_of Int.thy addsimps [not_neg_eq_ge_0, nat_number_of_def, neg_nat, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   186
				  nat_div_distrib RS sym, nat_0]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   187
qed "div_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   188
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   189
Addsimps [div_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   190
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   191
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   192
(** Remainder **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   193
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   194
(*Fails if z'<0: the LHS collapses to (nat z) but the RHS doesn't*)
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   195
Goal "[| (0::int) <= z;  0 <= z' |] ==> nat (z mod z') = nat z mod nat z'";
13183
c7290200b3f4 conversion of IntDiv.thy to Isar format
paulson
parents: 12613
diff changeset
   196
by (case_tac "z' = 0" 1 THEN asm_simp_tac (simpset() addsimps [DIVISION_BY_ZERO]) 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   197
 by (simp_tac (simpset() addsimps [numeral_0_eq_0, DIVISION_BY_ZERO_MOD]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   198
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset()));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   199
by (rename_tac "m m'" 1);
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   200
by (subgoal_tac "0 <= int m mod int m'" 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   201
 by (asm_full_simp_tac 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   202
     (simpset() addsimps [nat_less_iff, numeral_0_eq_0, pos_mod_sign]) 2);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   203
by (rtac (inj_int RS injD) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   204
by (Asm_simp_tac 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   205
by (res_inst_tac [("q", "int (m div m')")] quorem_mod 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   206
 by (Force_tac 2);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   207
by (asm_full_simp_tac 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   208
     (simpset() addsimps [nat_less_iff RS sym, quorem_def, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   209
		          numeral_0_eq_0, zadd_int, zmult_int]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   210
by (rtac (mod_div_equality RS sym RS trans) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   211
by (asm_simp_tac (simpset() addsimps add_ac@mult_ac) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   212
qed "nat_mod_distrib";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   213
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   214
Goal "(number_of v :: nat)  mod  number_of v' = \
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   215
\       (if neg (number_of v) then 0 \
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   216
\        else if neg (number_of v') then number_of v \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   217
\        else nat (number_of v mod number_of v'))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   218
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   219
    (simpset_of Int.thy addsimps [not_neg_eq_ge_0, nat_number_of_def, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   220
				  neg_nat, nat_0, DIVISION_BY_ZERO_MOD,
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   221
				  nat_mod_distrib RS sym]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   222
qed "mod_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   223
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   224
Addsimps [mod_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   225
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   226
structure NatAbstractNumeralsData =
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   227
  struct
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   228
  val dest_eq		= HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   229
  val is_numeral	= Bin_Simprocs.is_numeral
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   230
  val numeral_0_eq_0    = numeral_0_eq_0
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   231
  val numeral_1_eq_1    = numeral_1_eq_Suc_0
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   232
  val prove_conv   = Bin_Simprocs.prove_conv_nohyps "nat_abstract_numerals"
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   233
  fun norm_tac simps	= ALLGOALS (simp_tac (HOL_ss addsimps simps))
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   234
  val simplify_meta_eq  = Bin_Simprocs.simplify_meta_eq 
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   235
  end
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   236
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   237
structure NatAbstractNumerals = AbstractNumeralsFun (NatAbstractNumeralsData)
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   238
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   239
val nat_eval_numerals = 
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   240
  map Bin_Simprocs.prep_simproc
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   241
   [("nat_div_eval_numerals",
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   242
     Bin_Simprocs.prep_pats ["(Suc 0) div m"],
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   243
     NatAbstractNumerals.proc div_nat_number_of),
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   244
    ("nat_mod_eval_numerals",
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   245
     Bin_Simprocs.prep_pats ["(Suc 0) mod m"],
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   246
     NatAbstractNumerals.proc mod_nat_number_of)];
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   247
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   248
Addsimprocs nat_eval_numerals;
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   249
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   250
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   251
(*** Comparisons ***)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   252
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   253
(** Equals (=) **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   254
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   255
Goal "[| (0::int) <= z;  0 <= z' |] ==> (nat z = nat z') = (z=z')";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   256
by (auto_tac (claset() addSEs [nonneg_eq_int], simpset()));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   257
qed "eq_nat_nat_iff";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   258
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   259
(*"neg" is used in rewrite rules for binary comparisons*)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   260
Goal "((number_of v :: nat) = number_of v') = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   261
\     (if neg (number_of v) then (iszero (number_of v') | neg (number_of v')) \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   262
\      else if neg (number_of v') then iszero (number_of v) \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   263
\      else iszero (number_of (bin_add v (bin_minus v'))))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   264
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   265
    (simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   266
				  eq_nat_nat_iff, eq_number_of_eq, nat_0]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   267
by (simp_tac (simpset_of Int.thy addsimps [nat_eq_iff, nat_eq_iff2, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   268
					   iszero_def]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   269
by (simp_tac (simpset () addsimps [not_neg_eq_ge_0 RS sym]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   270
qed "eq_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   271
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   272
Addsimps [eq_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   273
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   274
(** Less-than (<) **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   275
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   276
(*"neg" is used in rewrite rules for binary comparisons*)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   277
Goal "((number_of v :: nat) < number_of v') = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   278
\        (if neg (number_of v) then neg (number_of (bin_minus v')) \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   279
\         else neg (number_of (bin_add v (bin_minus v'))))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   280
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   281
    (simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   282
				  nat_less_eq_zless, less_number_of_eq_neg,
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   283
				  nat_0]) 1);
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   284
by (simp_tac (simpset_of Int.thy addsimps [neg_eq_less_0, zminus_zless, 
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   285
				number_of_minus, zless_nat_eq_int_zless]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   286
qed "less_nat_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   287
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   288
Addsimps [less_nat_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   289
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   290
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   291
(** Less-than-or-equals (<=) **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   292
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   293
Goal "(number_of x <= (number_of y::nat)) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   294
\     (~ number_of y < (number_of x::nat))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   295
by (rtac (linorder_not_less RS sym) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   296
qed "le_nat_number_of_eq_not_less"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   297
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   298
Addsimps [le_nat_number_of_eq_not_less];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   299
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   300
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   301
(*Maps #n to n for n = 0, 1, 2*)
11018
71d624788ce2 added "numerals" theorems;
wenzelm
parents: 10960
diff changeset
   302
bind_thms ("numerals", [numeral_0_eq_0, numeral_1_eq_1, numeral_2_eq_2]);
71d624788ce2 added "numerals" theorems;
wenzelm
parents: 10960
diff changeset
   303
val numeral_ss = simpset() addsimps numerals;
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   304
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   305
(** Nat **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   306
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   307
Goal "0 < n ==> n = Suc(n - 1)";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   308
by (asm_full_simp_tac numeral_ss 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   309
qed "Suc_pred'";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   310
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   311
(*Expresses a natural number constant as the Suc of another one.
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   312
  NOT suitable for rewriting because n recurs in the condition.*)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   313
bind_thm ("expand_Suc", inst "n" "number_of ?v" Suc_pred');
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   314
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   315
(** Arith **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   316
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   317
Goal "Suc n = n + 1";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   318
by (asm_simp_tac numeral_ss 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   319
qed "Suc_eq_add_numeral_1";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   320
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   321
(* These two can be useful when m = number_of... *)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   322
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   323
Goal "(m::nat) + n = (if m=0 then n else Suc ((m - 1) + n))";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   324
by (case_tac "m" 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   325
by (ALLGOALS (asm_simp_tac numeral_ss));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   326
qed "add_eq_if";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   327
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   328
Goal "(m::nat) * n = (if m=0 then 0 else n + ((m - 1) * n))";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   329
by (case_tac "m" 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   330
by (ALLGOALS (asm_simp_tac numeral_ss));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   331
qed "mult_eq_if";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   332
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   333
Goal "(p ^ m :: nat) = (if m=0 then 1 else p * (p ^ (m - 1)))";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   334
by (case_tac "m" 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   335
by (ALLGOALS (asm_simp_tac numeral_ss));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   336
qed "power_eq_if";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   337
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   338
Goal "[| 0<n; 0<m |] ==> m - n < (m::nat)";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   339
by (asm_full_simp_tac (numeral_ss addsimps [diff_less]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   340
qed "diff_less'";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   341
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   342
Addsimps [inst "n" "number_of ?v" diff_less'];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   343
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   344
(** Power **)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   345
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   346
Goal "(p::nat) ^ 2 = p*p";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   347
by (simp_tac numeral_ss 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   348
qed "power_two";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   349
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   350
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   351
(*** Comparisons involving (0::nat) ***)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   352
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   353
Goal "(number_of v = (0::nat)) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   354
\     (if neg (number_of v) then True else iszero (number_of v))";
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   355
by (simp_tac (simpset() addsimps [numeral_0_eq_0 RS sym, iszero_0]) 1);
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   356
qed "eq_number_of_0"; 
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   357
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   358
Goal "((0::nat) = number_of v) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   359
\     (if neg (number_of v) then True else iszero (number_of v))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   360
by (rtac ([eq_sym_conv, eq_number_of_0] MRS trans) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   361
qed "eq_0_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   362
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   363
Goal "((0::nat) < number_of v) = neg (number_of (bin_minus v))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   364
by (simp_tac (simpset() addsimps [numeral_0_eq_0 RS sym]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   365
qed "less_0_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   366
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   367
(*Simplification already handles n<0, n<=0 and 0<=n.*)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   368
Addsimps [eq_number_of_0, eq_0_number_of, less_0_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   369
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   370
Goal "neg (number_of v) ==> number_of v = (0::nat)";
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   371
by (asm_simp_tac (simpset() addsimps [numeral_0_eq_0 RS sym, iszero_0]) 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   372
qed "neg_imp_number_of_eq_0";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   373
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   374
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   375
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   376
(*** Comparisons involving Suc ***)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   377
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   378
Goal "(number_of v = Suc n) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   379
\       (let pv = number_of (bin_pred v) in \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   380
\        if neg pv then False else nat pv = n)";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   381
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   382
    (simpset_of Int.thy addsimps
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   383
      [Let_def, neg_eq_less_0, linorder_not_less, number_of_pred,
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   384
       nat_number_of_def, zadd_0] @ zadd_ac) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   385
by (res_inst_tac [("x", "number_of v")] spec 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   386
by (auto_tac (claset(), simpset() addsimps [nat_eq_iff]));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   387
qed "eq_number_of_Suc";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   388
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   389
Goal "(Suc n = number_of v) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   390
\       (let pv = number_of (bin_pred v) in \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   391
\        if neg pv then False else nat pv = n)";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   392
by (rtac ([eq_sym_conv, eq_number_of_Suc] MRS trans) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   393
qed "Suc_eq_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   394
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   395
Goal "(number_of v < Suc n) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   396
\       (let pv = number_of (bin_pred v) in \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   397
\        if neg pv then True else nat pv < n)";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   398
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   399
    (simpset_of Int.thy addsimps
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   400
      [Let_def, neg_eq_less_0, linorder_not_less, number_of_pred,
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   401
       nat_number_of_def, zadd_0] @ zadd_ac) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   402
by (res_inst_tac [("x", "number_of v")] spec 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   403
by (auto_tac (claset(), simpset() addsimps [nat_less_iff]));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   404
qed "less_number_of_Suc";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   405
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   406
Goal "(Suc n < number_of v) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   407
\       (let pv = number_of (bin_pred v) in \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   408
\        if neg pv then False else n < nat pv)";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   409
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   410
    (simpset_of Int.thy addsimps
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   411
      [Let_def, neg_eq_less_0, linorder_not_less, number_of_pred,
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   412
       nat_number_of_def, zadd_0] @ zadd_ac) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   413
by (res_inst_tac [("x", "number_of v")] spec 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   414
by (auto_tac (claset(), simpset() addsimps [zless_nat_eq_int_zless]));
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   415
qed "less_Suc_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   416
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   417
Goal "(number_of v <= Suc n) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   418
\       (let pv = number_of (bin_pred v) in \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   419
\        if neg pv then True else nat pv <= n)";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   420
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   421
    (simpset () addsimps
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   422
      [Let_def, less_Suc_number_of, linorder_not_less RS sym]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   423
qed "le_number_of_Suc";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   424
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   425
Goal "(Suc n <= number_of v) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   426
\       (let pv = number_of (bin_pred v) in \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   427
\        if neg pv then False else n <= nat pv)";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   428
by (simp_tac
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   429
    (simpset () addsimps
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   430
      [Let_def, less_number_of_Suc, linorder_not_less RS sym]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   431
qed "le_Suc_number_of";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   432
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   433
Addsimps [eq_number_of_Suc, Suc_eq_number_of, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   434
	  less_number_of_Suc, less_Suc_number_of, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   435
	  le_number_of_Suc, le_Suc_number_of];
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   436
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   437
(* Push int(.) inwards: *)
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   438
Addsimps [zadd_int RS sym];
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   439
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   440
Goal "(m+m = n+n) = (m = (n::int))";
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   441
by Auto_tac;
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   442
val lemma1 = result();
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   443
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   444
Goal "m+m ~= (1::int) + n + n";
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   445
by Auto_tac;
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   446
by (dres_inst_tac [("f", "%x. x mod 2")] arg_cong 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   447
by (full_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   448
val lemma2 = result();
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   449
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   450
Goal "((number_of (v BIT x) ::int) = number_of (w BIT y)) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   451
\     (x=y & (((number_of v) ::int) = number_of w))"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   452
by (simp_tac (simpset_of Int.thy addsimps
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   453
	       [number_of_BIT, lemma1, lemma2, eq_commute]) 1); 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   454
qed "eq_number_of_BIT_BIT"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   455
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   456
Goal "((number_of (v BIT x) ::int) = number_of Pls) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   457
\     (x=False & (((number_of v) ::int) = number_of Pls))"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   458
by (simp_tac (simpset_of Int.thy addsimps
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   459
	       [number_of_BIT, number_of_Pls, eq_commute]) 1); 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   460
by (res_inst_tac [("x", "number_of v")] spec 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   461
by Safe_tac;
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   462
by (ALLGOALS Full_simp_tac);
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   463
by (dres_inst_tac [("f", "%x. x mod 2")] arg_cong 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   464
by (full_simp_tac (simpset() addsimps [zmod_zadd1_eq]) 1); 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   465
qed "eq_number_of_BIT_Pls"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   466
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   467
Goal "((number_of (v BIT x) ::int) = number_of Min) = \
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   468
\     (x=True & (((number_of v) ::int) = number_of Min))"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   469
by (simp_tac (simpset_of Int.thy addsimps
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   470
	       [number_of_BIT, number_of_Min, eq_commute]) 1); 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   471
by (res_inst_tac [("x", "number_of v")] spec 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   472
by Auto_tac;
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   473
by (dres_inst_tac [("f", "%x. x mod 2")] arg_cong 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   474
by Auto_tac;
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   475
qed "eq_number_of_BIT_Min"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   476
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   477
Goal "(number_of Pls ::int) ~= number_of Min"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   478
by Auto_tac;
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   479
qed "eq_number_of_Pls_Min"; 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   480
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   481
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   482
(*** Further lemmas about "nat" ***)
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   483
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   484
Goal "nat (abs (w * z)) = nat (abs w) * nat (abs z)";
11868
56db9f3a6b3e Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents: 11704
diff changeset
   485
by (case_tac "z=0 | w=0" 1);
10574
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   486
by Auto_tac;  
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   487
by (simp_tac (simpset() addsimps [zabs_def, nat_mult_distrib RS sym, 
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   488
                          nat_mult_distrib_neg RS sym, zmult_less_0_iff]) 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   489
by (arith_tac 1);
8f98f0301d67 Linear arithmetic now copes with mixed nat/int formulae.
nipkow
parents:
diff changeset
   490
qed "nat_abs_mult_distrib";
10754
9bc30e51144c now #16*(x+y) distributes for nat just as for other numeric types
paulson
parents: 10710
diff changeset
   491
9bc30e51144c now #16*(x+y) distributes for nat just as for other numeric types
paulson
parents: 10710
diff changeset
   492
(*Distributive laws for literals*)
9bc30e51144c now #16*(x+y) distributes for nat just as for other numeric types
paulson
parents: 10710
diff changeset
   493
Addsimps (map (inst "k" "number_of ?v")
9bc30e51144c now #16*(x+y) distributes for nat just as for other numeric types
paulson
parents: 10710
diff changeset
   494
	  [add_mult_distrib, add_mult_distrib2,
9bc30e51144c now #16*(x+y) distributes for nat just as for other numeric types
paulson
parents: 10710
diff changeset
   495
	   diff_mult_distrib, diff_mult_distrib2]);
9bc30e51144c now #16*(x+y) distributes for nat just as for other numeric types
paulson
parents: 10710
diff changeset
   496
12613
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   497
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   498
(*** Literal arithmetic involving powers, type nat ***)
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   499
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   500
Goal "(0::int) <= z ==> nat (z^n) = nat z ^ n";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   501
by (induct_tac "n" 1); 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   502
by (ALLGOALS (asm_simp_tac (simpset() addsimps [nat_mult_distrib])));
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   503
qed "nat_power_eq";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   504
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   505
Goal "(number_of v :: nat) ^ n = \
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   506
\      (if neg (number_of v) then 0^n else nat ((number_of v :: int) ^ n))";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   507
by (simp_tac
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   508
    (simpset_of Int.thy addsimps [neg_nat, not_neg_eq_ge_0, nat_number_of_def, 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   509
				  nat_power_eq]) 1);
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   510
qed "power_nat_number_of";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   511
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   512
Addsimps [inst "n" "number_of ?w" power_nat_number_of];
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   513
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   514
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   515
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   516
(*** Literal arithmetic involving powers, type int ***)
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   517
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   518
Goal "(z::int) ^ (2*a) = (z^a)^2";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   519
by (simp_tac (simpset() addsimps [zpower_zpower, mult_commute]) 1); 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   520
qed "zpower_even";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   521
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   522
Goal "(p::int) ^ 2 = p*p"; 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   523
by (simp_tac numeral_ss 1);
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   524
qed "zpower_two";  
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   525
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   526
Goal "(z::int) ^ (2*a + 1) = z * (z^a)^2";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   527
by (simp_tac (simpset() addsimps [zpower_even, zpower_zadd_distrib]) 1); 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   528
qed "zpower_odd";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   529
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   530
Goal "(z::int) ^ number_of (w BIT False) = \
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   531
\     (let w = z ^ (number_of w) in  w*w)";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   532
by (simp_tac (simpset_of Int.thy addsimps [nat_number_of_def,
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   533
        number_of_BIT, Let_def]) 1);
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   534
by (res_inst_tac [("x","number_of w")] spec 1 THEN Clarify_tac 1); 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   535
by (case_tac "(0::int) <= x" 1);
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   536
by (auto_tac (claset(), 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   537
     simpset() addsimps [nat_mult_distrib, zpower_even, zpower_two])); 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   538
qed "zpower_number_of_even";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   539
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   540
Goal "(z::int) ^ number_of (w BIT True) = \
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   541
\         (if (0::int) <= number_of w                   \
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   542
\          then (let w = z ^ (number_of w) in  z*w*w)   \
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   543
\          else 1)";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   544
by (simp_tac (simpset_of Int.thy addsimps [nat_number_of_def,
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   545
        number_of_BIT, Let_def]) 1);
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   546
by (res_inst_tac [("x","number_of w")] spec 1 THEN Clarify_tac 1); 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   547
by (case_tac "(0::int) <= x" 1);
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   548
by (auto_tac (claset(), 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   549
              simpset() addsimps [nat_add_distrib, nat_mult_distrib, 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   550
                                  zpower_even, zpower_two])); 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   551
qed "zpower_number_of_odd";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   552
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   553
Addsimps (map (inst "z" "number_of ?v")
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   554
              [zpower_number_of_even, zpower_number_of_odd]);
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12196
diff changeset
   555